diff -r a0334d7fb0ab -r bdca9f765ee4 src/HOL/Hoare_Parallel/RG_Syntax.thy --- a/src/HOL/Hoare_Parallel/RG_Syntax.thy Thu Feb 11 09:14:34 2010 +0100 +++ b/src/HOL/Hoare_Parallel/RG_Syntax.thy Thu Feb 11 13:54:53 2010 +0100 @@ -4,10 +4,13 @@ imports RG_Hoare Quote_Antiquote begin +abbreviation Skip :: "'a com" ("SKIP") + where "SKIP \ Basic id" + +notation Seq ("(_;;/ _)" [60,61] 60) + syntax "_Assign" :: "idt \ 'b \ 'a com" ("(\_ :=/ _)" [70, 65] 61) - "_skip" :: "'a com" ("SKIP") - "_Seq" :: "'a com \ 'a com \ 'a com" ("(_;;/ _)" [60,61] 60) "_Cond" :: "'a bexp \ 'a com \ 'a com \ 'a com" ("(0IF _/ THEN _/ ELSE _/FI)" [0, 0, 0] 61) "_Cond2" :: "'a bexp \ 'a com \ 'a com" ("(0IF _ THEN _ FI)" [0,0] 56) "_While" :: "'a bexp \ 'a com \ 'a com" ("(0WHILE _ /DO _ /OD)" [0, 0] 61) @@ -16,9 +19,7 @@ "_Wait" :: "'a bexp \ 'a com" ("(0WAIT _ END)" 61) translations - "\\x := a" \ "Basic \\\(_update_name x (\_. a))\" - "SKIP" \ "Basic id" - "c1;; c2" \ "Seq c1 c2" + "\x := a" \ "CONST Basic \\(_update_name x (\_. a))\" "IF b THEN c1 ELSE c2 FI" \ "Cond .{b}. c1 c2" "IF b THEN c FI" \ "IF b THEN c ELSE SKIP FI" "WHILE b DO c OD" \ "While .{b}. c" @@ -52,8 +53,8 @@ "_after" :: "id \ 'a" ("\_") translations - "\x" \ "x \fst" - "\x" \ "x \snd" + "\x" \ "x \CONST fst" + "\x" \ "x \CONST snd" print_translation {* let @@ -63,7 +64,7 @@ val assert_tr' = quote_tr' (Syntax.const "_Assert"); - fun bexp_tr' name ((Const ("Collect", _) $ t) :: ts) = + fun bexp_tr' name ((Const (@{const_syntax Collect}, _) $ t) :: ts) = quote_tr' (Syntax.const name) (t :: ts) | bexp_tr' _ _ = raise Match; @@ -78,8 +79,10 @@ | update_name_tr' (Const x) = Const (upd_tr' x) | update_name_tr' _ = raise Match; - fun K_tr' (Abs (_,_,t)) = if null (loose_bnos t) then t else raise Match - | K_tr' (Abs (_,_,Abs (_,_,t)$Bound 0)) = if null (loose_bnos t) then t else raise Match + fun K_tr' (Abs (_, _, t)) = + if null (loose_bnos t) then t else raise Match + | K_tr' (Abs (_, _, Abs (_, _, t) $ Bound 0)) = + if null (loose_bnos t) then t else raise Match | K_tr' _ = raise Match; fun assign_tr' (Abs (x, _, f $ k $ Bound 0) :: ts) = @@ -87,8 +90,10 @@ (Abs (x, dummyT, K_tr' k) :: ts) | assign_tr' _ = raise Match; in - [("Collect", assert_tr'), ("Basic", assign_tr'), - ("Cond", bexp_tr' "_Cond"), ("While", bexp_tr' "_While_inv")] + [(@{const_syntax Collect}, assert_tr'), + (@{const_syntax Basic}, assign_tr'), + (@{const_syntax Cond}, bexp_tr' "_Cond"), + (@{const_syntax While}, bexp_tr' "_While_inv")] end *}