diff -r 66b62cbeaab3 -r 6539627881e8 src/HOL/NanoJava/AxSem.thy --- a/src/HOL/NanoJava/AxSem.thy Mon Sep 10 13:57:57 2001 +0200 +++ b/src/HOL/NanoJava/AxSem.thy Mon Sep 10 17:35:22 2001 +0200 @@ -38,7 +38,7 @@ translations "A |\ C" \ "(A,C) \ hoare" "A \ {P}c{Q}" \ "A |\ {(P,c,Q)}" "A |\\<^sub>e t" \ "(A,t) \ ehoare" - "A |\\<^sub>e (P,e,Q)" \ "(A,P,e,Q) \ ehoare" (** shouldn't be necess.**) + "A |\\<^sub>e (P,e,Q)" \ "(A,P,e,Q) \ ehoare" (* shouldn't be necessary *) "A \\<^sub>e {P}e{Q}" \ "A |\\<^sub>e (P,e,Q)" @@ -77,7 +77,7 @@ Call: "[| A |-e {P} e1 {Q}; \a. A |-e {Q a} e2 {R a}; \a p l. A |- {\s'. \s. R a p s \ l = s \ - s' = lupd(This\a)(lupd(Param\p)(reset_locs s))} + s' = lupd(This\a)(lupd(Par\p)(reset_locs s))} Meth (C,m) {\s. S (s) (set_locs l s)} |] ==> A |-e {P} {C}e1..m(e2) {S}" @@ -86,13 +86,13 @@ Impl (D,m) {Q} ==> A |- {P} Meth (C,m) {Q}" - (*\z instead of \z in the conclusion and - z restricted to type state due to limitations of the inductive package *) + --{*\z instead of \z in the conclusion and + z restricted to type state due to limitations of the inductive package *} Impl: "\z::state. A\ (\z. (\Cm. (P z Cm, Impl Cm, Q z Cm))`Ms) ||- (\Cm. (P z Cm, body Cm, Q z Cm))`Ms ==> A ||- (\Cm. (P z Cm, Impl Cm, Q z Cm))`Ms" -(* structural rules *) +--{* structural rules *} Asm: " a \ A ==> A ||- {a}" @@ -104,7 +104,7 @@ \s t. (\z. P' z s --> Q' z t) --> (P s --> Q t) |] ==> A |- {P} c {Q }" - (* z restricted to type state due to limitations of the inductive package *) + --{* z restricted to type state due to limitations of the inductive package *} eConseq:"[| \z::state. A |-e {P' z} c {Q' z}; \s v t. (\z. P' z s --> Q' z v t) --> (P s --> Q v t) |] ==> A |-e {P} c {Q }"