diff -r a7790c8e3c14 -r e4a3f49eb924 src/HOL/HoareParallel/RG_Syntax.thy --- a/src/HOL/HoareParallel/RG_Syntax.thy Fri Apr 20 17:58:24 2007 +0200 +++ b/src/HOL/HoareParallel/RG_Syntax.thy Fri Apr 20 17:58:25 2007 +0200 @@ -16,7 +16,7 @@ "_Wait" :: "'a bexp \ 'a com" ("(0WAIT _ END)" 61) translations - "\\x := a" \ "Basic \\\(_update_name x (CONST K_record a))\" + "\\x := a" \ "Basic \\\(_update_name x (K_record a))\" "SKIP" \ "Basic id" "c1;; c2" \ "Seq c1 c2" "IF b THEN c1 ELSE c2 FI" \ "Cond .{b}. c1 c2" @@ -78,7 +78,7 @@ | update_name_tr' (Const x) = Const (upd_tr' x) | update_name_tr' _ = raise Match; - fun assign_tr' (Abs (x, _, f $ (Const ("K_record",_)$t) $ Bound 0) :: ts) = + fun assign_tr' (Abs (x, _, f $ (Const (@{const_syntax "K_record"}, _)$t) $ Bound 0) :: ts) = quote_tr' (Syntax.const "_Assign" $ update_name_tr' f) (Abs (x, dummyT, t) :: ts) | assign_tr' _ = raise Match;