src/HOL/HoareParallel/RG_Syntax.thy
changeset 13020 791e3b4c4039
child 13022 b115b305612f
equal deleted inserted replaced
13019:98f0a09a33c3 13020:791e3b4c4039
       
     1 
       
     2 header {* \section{Concrete Syntax} *}
       
     3 
       
     4 theory RG_Syntax = Quote_Antiquote + RG_Hoare:
       
     5 
       
     6 syntax
       
     7   "_Assign"    :: "idt \<Rightarrow> 'b \<Rightarrow> 'a com"                     ("(\<acute>_ :=/ _)" [70, 65] 61)
       
     8   "_skip"      :: "'a com"                                  ("SKIP")
       
     9   "_Seq"       :: "'a com \<Rightarrow> 'a com \<Rightarrow> 'a com"              ("(_;;/ _)" [60,61] 60)
       
    10   "_Cond"      :: "'a bexp \<Rightarrow> 'a com \<Rightarrow> 'a com \<Rightarrow> 'a com"   ("(0IF _/ THEN _/ ELSE _/FI)" [0, 0, 0] 61)
       
    11   "_Cond2"     :: "'a bexp \<Rightarrow> 'a com \<Rightarrow> 'a com"             ("(0IF _ THEN _ FI)" [0,0] 56)
       
    12   "_While"     :: "'a bexp \<Rightarrow> 'a com \<Rightarrow> 'a com"             ("(0WHILE _ /DO _ /OD)"  [0, 0] 61)
       
    13   "_Await"     :: "'a bexp \<Rightarrow> 'a com \<Rightarrow> 'a com"             ("(0AWAIT _ /THEN /_ /END)"  [0,0] 61)
       
    14   "_Atom"      :: "'a com \<Rightarrow> 'a com"                        ("(\<langle>_\<rangle>)" 61)
       
    15   "_Wait"      :: "'a bexp \<Rightarrow> 'a com"                       ("(0WAIT _ END)" 61)
       
    16 
       
    17 translations
       
    18   "\<acute>\<spacespace>x := a" \<rightharpoonup> "Basic \<guillemotleft>\<acute>\<spacespace>(_update_name x a)\<guillemotright>"
       
    19   "SKIP" \<rightleftharpoons> "Basic id"
       
    20   "c1;; c2" \<rightleftharpoons> "Seq c1 c2"
       
    21   "IF b THEN c1 ELSE c2 FI" \<rightharpoonup> "Cond .{b}. c1 c2"
       
    22   "IF b THEN c FI" \<rightleftharpoons> "IF b THEN c ELSE SKIP FI"
       
    23   "WHILE b DO c OD" \<rightharpoonup> "While .{b}. c"
       
    24   "AWAIT b THEN c END" \<rightleftharpoons> "Await .{b}. c"
       
    25   "\<langle>c\<rangle>" \<rightleftharpoons> "AWAIT True THEN c END"
       
    26   "WAIT b END" \<rightleftharpoons> "AWAIT b THEN SKIP END"
       
    27 
       
    28 nonterminals
       
    29   prgs
       
    30 
       
    31 syntax
       
    32   "_PAR"        :: "prgs \<Rightarrow> 'a"              ("COBEGIN//_//COEND" 60)
       
    33   "_prg"        :: "'a \<Rightarrow> prgs"              ("_" 57)
       
    34   "_prgs"       :: "['a, prgs] \<Rightarrow> prgs"      ("_//\<parallel>//_" [60,57] 57)
       
    35 
       
    36 translations
       
    37   "_prg a" \<rightharpoonup> "[a]"
       
    38   "_prgs a ps" \<rightharpoonup> "a # ps"
       
    39   "_PAR ps" \<rightharpoonup> "ps"
       
    40 
       
    41 syntax
       
    42   "_prg_scheme" :: "['a, 'a, 'a, 'a] \<Rightarrow> prgs"  ("SCHEME [_ \<le> _ < _] _" [0,0,0,60] 57)
       
    43 
       
    44 translations
       
    45   "_prg_scheme j i k c" \<rightleftharpoons> "(map (\<lambda>i. c) [j..k(])"
       
    46 
       
    47 text {* Translations for variables before and after a transition: *}
       
    48 
       
    49 syntax 
       
    50   "_before" :: "id \<Rightarrow> 'a" ("\<ordmasculine>_")
       
    51   "_after"  :: "id \<Rightarrow> 'a" ("\<ordfeminine>_")
       
    52  
       
    53 translations
       
    54   "\<ordmasculine>x" \<rightleftharpoons> "x \<acute>fst"
       
    55   "\<ordfeminine>x" \<rightleftharpoons> "x \<acute>snd"
       
    56 
       
    57 print_translation {*
       
    58   let
       
    59     fun quote_tr' f (t :: ts) =
       
    60           Term.list_comb (f $ Syntax.quote_tr' "_antiquote" t, ts)
       
    61       | quote_tr' _ _ = raise Match;
       
    62 
       
    63     val assert_tr' = quote_tr' (Syntax.const "_Assert");
       
    64 
       
    65     fun bexp_tr' name ((Const ("Collect", _) $ t) :: ts) =
       
    66           quote_tr' (Syntax.const name) (t :: ts)
       
    67       | bexp_tr' _ _ = raise Match;
       
    68 
       
    69     fun upd_tr' (x_upd, T) =
       
    70       (case try (unsuffix RecordPackage.updateN) x_upd of
       
    71         Some x => (x, if T = dummyT then T else Term.domain_type T)
       
    72       | None => raise Match);
       
    73 
       
    74     fun update_name_tr' (Free x) = Free (upd_tr' x)
       
    75       | update_name_tr' ((c as Const ("_free", _)) $ Free x) =
       
    76           c $ Free (upd_tr' x)
       
    77       | update_name_tr' (Const x) = Const (upd_tr' x)
       
    78       | update_name_tr' _ = raise Match;
       
    79 
       
    80     fun assign_tr' (Abs (x, _, f $ t $ Bound 0) :: ts) =
       
    81           quote_tr' (Syntax.const "_Assign" $ update_name_tr' f)
       
    82             (Abs (x, dummyT, t) :: ts)
       
    83       | assign_tr' _ = raise Match;
       
    84   in
       
    85     [("Collect", assert_tr'), ("Basic", assign_tr'),
       
    86       ("Cond", bexp_tr' "_Cond"), ("While", bexp_tr' "_While_inv")]
       
    87   end
       
    88 *}
       
    89 
       
    90 end