diff -r 35094c8fd8bf -r a073cb249a06 src/HOL/Hoare_Parallel/RG_Syntax.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Hoare_Parallel/RG_Syntax.thy Mon Sep 21 10:58:25 2009 +0200 @@ -0,0 +1,95 @@ +header {* \section{Concrete Syntax} *} + +theory RG_Syntax +imports RG_Hoare Quote_Antiquote +begin + +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) + "_Await" :: "'a bexp \ 'a com \ 'a com" ("(0AWAIT _ /THEN /_ /END)" [0,0] 61) + "_Atom" :: "'a com \ 'a com" ("(\_\)" 61) + "_Wait" :: "'a bexp \ 'a com" ("(0WAIT _ END)" 61) + +translations + "\\x := a" \ "Basic \\\(_update_name x (\_. a))\" + "SKIP" \ "Basic id" + "c1;; c2" \ "Seq c1 c2" + "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" + "AWAIT b THEN c END" \ "Await .{b}. c" + "\c\" \ "AWAIT True THEN c END" + "WAIT b END" \ "AWAIT b THEN SKIP END" + +nonterminals + prgs + +syntax + "_PAR" :: "prgs \ 'a" ("COBEGIN//_//COEND" 60) + "_prg" :: "'a \ prgs" ("_" 57) + "_prgs" :: "['a, prgs] \ prgs" ("_//\//_" [60,57] 57) + +translations + "_prg a" \ "[a]" + "_prgs a ps" \ "a # ps" + "_PAR ps" \ "ps" + +syntax + "_prg_scheme" :: "['a, 'a, 'a, 'a] \ prgs" ("SCHEME [_ \ _ < _] _" [0,0,0,60] 57) + +translations + "_prg_scheme j i k c" \ "(map (\i. c) [j.. 'a" ("\_") + "_after" :: "id \ 'a" ("\_") + +translations + "\x" \ "x \fst" + "\x" \ "x \snd" + +print_translation {* + let + fun quote_tr' f (t :: ts) = + Term.list_comb (f $ Syntax.quote_tr' "_antiquote" t, ts) + | quote_tr' _ _ = raise Match; + + val assert_tr' = quote_tr' (Syntax.const "_Assert"); + + fun bexp_tr' name ((Const ("Collect", _) $ t) :: ts) = + quote_tr' (Syntax.const name) (t :: ts) + | bexp_tr' _ _ = raise Match; + + fun upd_tr' (x_upd, T) = + (case try (unsuffix Record.updateN) x_upd of + SOME x => (x, if T = dummyT then T else Term.domain_type T) + | NONE => raise Match); + + fun update_name_tr' (Free x) = Free (upd_tr' x) + | update_name_tr' ((c as Const ("_free", _)) $ Free x) = + c $ Free (upd_tr' x) + | 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 + | K_tr' _ = raise Match; + + fun assign_tr' (Abs (x, _, f $ k $ Bound 0) :: ts) = + quote_tr' (Syntax.const "_Assign" $ update_name_tr' f) + (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")] + end +*} + +end \ No newline at end of file