# HG changeset patch # User wenzelm # Date 1265892893 -3600 # Node ID bdca9f765ee4d6181b1129ff8caee54c24741f68 # Parent a0334d7fb0abbdb18c30924cba78c349a1fab75f modernized translations; diff -r a0334d7fb0ab -r bdca9f765ee4 src/HOL/Hoare_Parallel/OG_Syntax.thy --- a/src/HOL/Hoare_Parallel/OG_Syntax.thy Thu Feb 11 09:14:34 2010 +0100 +++ b/src/HOL/Hoare_Parallel/OG_Syntax.thy Thu Feb 11 13:54:53 2010 +0100 @@ -5,18 +5,25 @@ text{* Syntax for commands and for assertions and boolean expressions in commands @{text com} and annotated commands @{text ann_com}. *} +abbreviation Skip :: "'a com" ("SKIP" 63) + where "SKIP \ Basic id" + +abbreviation AnnSkip :: "'a assn \ 'a ann_com" ("_//SKIP" [90] 63) + where "r SKIP \ AnnBasic r id" + +notation + Seq ("_,,/ _" [55, 56] 55) and + AnnSeq ("_;;/ _" [60,61] 60) + syntax "_Assign" :: "idt \ 'b \ 'a com" ("(\_ :=/ _)" [70, 65] 61) "_AnnAssign" :: "'a assn \ idt \ 'b \ 'a com" ("(_ \_ :=/ _)" [90,70,65] 61) translations - "\\x := a" \ "Basic \\\(_update_name x (\_. a))\" - "r \\x := a" \ "AnnBasic r \\\(_update_name x (\_. a))\" + "\x := a" \ "CONST Basic \\(_update_name x (\_. a))\" + "r \x := a" \ "CONST AnnBasic r \\(_update_name x (\_. a))\" syntax - "_AnnSkip" :: "'a assn \ 'a ann_com" ("_//SKIP" [90] 63) - "_AnnSeq" :: "'a ann_com \ 'a ann_com \ 'a ann_com" ("_;;/ _" [60,61] 60) - "_AnnCond1" :: "'a assn \ 'a bexp \ 'a ann_com \ 'a ann_com \ 'a ann_com" ("_ //IF _ /THEN _ /ELSE _ /FI" [90,0,0,0] 61) "_AnnCond2" :: "'a assn \ 'a bexp \ 'a ann_com \ 'a ann_com" @@ -28,8 +35,6 @@ "_AnnAtom" :: "'a assn \ 'a com \ 'a ann_com" ("_//\_\" [90,0] 61) "_AnnWait" :: "'a assn \ 'a bexp \ 'a ann_com" ("_//WAIT _ END" [90,0] 61) - "_Skip" :: "'a com" ("SKIP" 63) - "_Seq" :: "'a com \ 'a com \ 'a com" ("_,,/ _" [55, 56] 55) "_Cond" :: "'a bexp \ 'a com \ 'a com \ 'a com" ("(0IF _/ THEN _/ ELSE _/ FI)" [0, 0, 0] 61) "_Cond2" :: "'a bexp \ 'a com \ 'a com" ("IF _ THEN _ FI" [0,0] 56) @@ -39,20 +44,15 @@ ("(0WHILE _ //DO _ /OD)" [0, 0] 61) translations - "SKIP" \ "Basic id" - "c_1,, c_2" \ "Seq c_1 c_2" - - "IF b THEN c1 ELSE c2 FI" \ "Cond .{b}. c1 c2" + "IF b THEN c1 ELSE c2 FI" \ "CONST Cond .{b}. c1 c2" "IF b THEN c FI" \ "IF b THEN c ELSE SKIP FI" "WHILE b INV i DO c OD" \ "While .{b}. i c" "WHILE b DO c OD" \ "WHILE b INV CONST undefined DO c OD" - "r SKIP" \ "AnnBasic r id" - "c_1;; c_2" \ "AnnSeq c_1 c_2" - "r IF b THEN c1 ELSE c2 FI" \ "AnnCond1 r .{b}. c1 c2" - "r IF b THEN c FI" \ "AnnCond2 r .{b}. c" - "r WHILE b INV i DO c OD" \ "AnnWhile r .{b}. i c" - "r AWAIT b THEN c END" \ "AnnAwait r .{b}. c" + "r IF b THEN c1 ELSE c2 FI" \ "CONST AnnCond1 r .{b}. c1 c2" + "r IF b THEN c FI" \ "CONST AnnCond2 r .{b}. c" + "r WHILE b INV i DO c OD" \ "CONST AnnWhile r .{b}. i c" + "r AWAIT b THEN c END" \ "CONST AnnAwait r .{b}. c" "r \c\" \ "r AWAIT True THEN c END" "r WAIT b END" \ "r AWAIT b THEN SKIP END" @@ -68,11 +68,11 @@ ("SCHEME [_ \ _ < _] _// _" [0,0,0,60, 90] 57) translations - "_prg c q" \ "[(Some c, q)]" - "_prgs c q ps" \ "(Some c, q) # ps" - "_PAR ps" \ "Parallel ps" + "_prg c q" \ "[(CONST Some c, q)]" + "_prgs c q ps" \ "(CONST Some c, q) # ps" + "_PAR ps" \ "CONST Parallel ps" - "_prg_scheme j i k c q" \ "CONST map (\i. (Some c, q)) [j.. "CONST map (\i. (CONST Some c, q)) [j.. 'a) \ nat \ ('a ann_com_op \ 'a) @@ -101,8 +101,8 @@ SEM :: "'a com \ 'a set \ 'a set" "SEM c S \ \sem c ` S " -syntax "_Omega" :: "'a com" ("\" 63) -translations "\" \ "While CONST UNIV CONST UNIV (Basic id)" +abbreviation Omega :: "'a com" ("\" 63) + where "\ \ While UNIV UNIV (Basic id)" consts fwhile :: "'a bexp \ 'a com \ nat \ 'a com" primrec diff -r a0334d7fb0ab -r bdca9f765ee4 src/HOL/Hoare_Parallel/Quote_Antiquote.thy --- a/src/HOL/Hoare_Parallel/Quote_Antiquote.thy Thu Feb 11 09:14:34 2010 +0100 +++ b/src/HOL/Hoare_Parallel/Quote_Antiquote.thy Thu Feb 11 13:54:53 2010 +0100 @@ -12,7 +12,7 @@ "_Assert" :: "'a \ 'a set" ("(\_\)" [0] 1000) translations - ".{b}." \ "Collect \b\" + ".{b}." \ "CONST Collect \b\" parse_translation {* let 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 *}