src/HOL/IMP/Hoare.thy
author nipkow
Wed, 08 Mar 1995 10:25:50 +0100
changeset 939 534955033ed2
parent 938 621be7ec81d7
child 972 e61b058d58d2
permissions -rw-r--r--
Added pretty-printing coments
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
936
a6d7b4084761 Hoare logic
nipkow
parents:
diff changeset
     1
(*  Title: 	HOL/IMP/Hoare.thy
938
621be7ec81d7 *** empty log message ***
nipkow
parents: 937
diff changeset
     2
    ID:         $Id$
936
a6d7b4084761 Hoare logic
nipkow
parents:
diff changeset
     3
    Author: 	Tobias Nipkow
a6d7b4084761 Hoare logic
nipkow
parents:
diff changeset
     4
    Copyright   1995 TUM
a6d7b4084761 Hoare logic
nipkow
parents:
diff changeset
     5
a6d7b4084761 Hoare logic
nipkow
parents:
diff changeset
     6
Semantic embedding of Hoare logic
a6d7b4084761 Hoare logic
nipkow
parents:
diff changeset
     7
*)
a6d7b4084761 Hoare logic
nipkow
parents:
diff changeset
     8
a6d7b4084761 Hoare logic
nipkow
parents:
diff changeset
     9
Hoare = Denotation +
a6d7b4084761 Hoare logic
nipkow
parents:
diff changeset
    10
consts
939
534955033ed2 Added pretty-printing coments
nipkow
parents: 938
diff changeset
    11
  spec :: "[state=>bool,com,state=>bool] => bool"
534955033ed2 Added pretty-printing coments
nipkow
parents: 938
diff changeset
    12
(* syntax "@spec" :: "[bool,com,bool] => bool" *)
534955033ed2 Added pretty-printing coments
nipkow
parents: 938
diff changeset
    13
          ("{{(1_)}}/ (_)/ {{(1_)}}" 10)
936
a6d7b4084761 Hoare logic
nipkow
parents:
diff changeset
    14
defs
a6d7b4084761 Hoare logic
nipkow
parents:
diff changeset
    15
  spec_def "spec P c Q == !s t. <s,t> : C(c) --> P s --> Q t"
a6d7b4084761 Hoare logic
nipkow
parents:
diff changeset
    16
end
939
534955033ed2 Added pretty-printing coments
nipkow
parents: 938
diff changeset
    17
(* Pretty-printing of assertions.
534955033ed2 Added pretty-printing coments
nipkow
parents: 938
diff changeset
    18
   Not very helpful as long as programs are not pretty-printed.
534955033ed2 Added pretty-printing coments
nipkow
parents: 938
diff changeset
    19
ML
534955033ed2 Added pretty-printing coments
nipkow
parents: 938
diff changeset
    20
534955033ed2 Added pretty-printing coments
nipkow
parents: 938
diff changeset
    21
local open Syntax
534955033ed2 Added pretty-printing coments
nipkow
parents: 938
diff changeset
    22
534955033ed2 Added pretty-printing coments
nipkow
parents: 938
diff changeset
    23
fun is_loc a = let val ch = hd(explode a)
534955033ed2 Added pretty-printing coments
nipkow
parents: 938
diff changeset
    24
               in ord "A" <= ord ch andalso ord ch <= ord "Z" end;
534955033ed2 Added pretty-printing coments
nipkow
parents: 938
diff changeset
    25
534955033ed2 Added pretty-printing coments
nipkow
parents: 938
diff changeset
    26
fun tr(s$t,i) = tr(s,i)$tr(t,i)
534955033ed2 Added pretty-printing coments
nipkow
parents: 938
diff changeset
    27
  | tr(Abs(x,T,u),i) = Abs(x,T,tr(u,i+1))
534955033ed2 Added pretty-printing coments
nipkow
parents: 938
diff changeset
    28
  | tr(t as Free(a,T),i) = if is_loc a then Bound(i) $ free(a) else t
534955033ed2 Added pretty-printing coments
nipkow
parents: 938
diff changeset
    29
  | tr(t,_) = t;
534955033ed2 Added pretty-printing coments
nipkow
parents: 938
diff changeset
    30
534955033ed2 Added pretty-printing coments
nipkow
parents: 938
diff changeset
    31
fun cond_tr(p) = Abs("",dummyT,tr(p,0))
534955033ed2 Added pretty-printing coments
nipkow
parents: 938
diff changeset
    32
534955033ed2 Added pretty-printing coments
nipkow
parents: 938
diff changeset
    33
fun spec_tr[p,c,q] = const"spec" $ cond_tr p $ c $ cond_tr q;
534955033ed2 Added pretty-printing coments
nipkow
parents: 938
diff changeset
    34
534955033ed2 Added pretty-printing coments
nipkow
parents: 938
diff changeset
    35
fun tr'(t as (Bound j $ (u as Free(a,_))),i) = if i=j then u else t
534955033ed2 Added pretty-printing coments
nipkow
parents: 938
diff changeset
    36
  | tr'(s$t,i) = tr'(s,i)$tr'(t,i)
534955033ed2 Added pretty-printing coments
nipkow
parents: 938
diff changeset
    37
  | tr'(Abs(x,T,u),i) = Abs(x,T,tr'(u,i+1))
534955033ed2 Added pretty-printing coments
nipkow
parents: 938
diff changeset
    38
  | tr'(t,_) = t;
534955033ed2 Added pretty-printing coments
nipkow
parents: 938
diff changeset
    39
534955033ed2 Added pretty-printing coments
nipkow
parents: 938
diff changeset
    40
fun spec_tr'[Abs(_,_,p),c,Abs(_,_,q)] =
534955033ed2 Added pretty-printing coments
nipkow
parents: 938
diff changeset
    41
  const"@spec" $ tr'(p,0) $ c $ tr'(q,0);
534955033ed2 Added pretty-printing coments
nipkow
parents: 938
diff changeset
    42
534955033ed2 Added pretty-printing coments
nipkow
parents: 938
diff changeset
    43
in
534955033ed2 Added pretty-printing coments
nipkow
parents: 938
diff changeset
    44
534955033ed2 Added pretty-printing coments
nipkow
parents: 938
diff changeset
    45
val parse_translation = [("@spec", spec_tr)];
534955033ed2 Added pretty-printing coments
nipkow
parents: 938
diff changeset
    46
val print_translation = [("spec", spec_tr')];
534955033ed2 Added pretty-printing coments
nipkow
parents: 938
diff changeset
    47
534955033ed2 Added pretty-printing coments
nipkow
parents: 938
diff changeset
    48
end
534955033ed2 Added pretty-printing coments
nipkow
parents: 938
diff changeset
    49
*)