Theory RG_Syntax

section ‹Concrete Syntax›

theory RG_Syntax
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)
  "_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"  "CONST Basic «´(_update_name x (λ_. a))»"
  "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 DO c OD"  "CONST While b c"
  "AWAIT b THEN c END"  "CONST Await b c"
  "c"  "AWAIT CONST True THEN c END"
  "WAIT b END"  "AWAIT b THEN SKIP END"

nonterminal 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"  "(CONST map (λi. c) [j..<k])"

text ‹Translations for variables before and after a transition:›

syntax
  "_before" :: "id  'a" ("º_")
  "_after"  :: "id  'a" ("ª_")

translations
  "ºx"  "x ´CONST fst"
  "ªx"  "x ´CONST snd"

print_translation let
    fun quote_tr' f (t :: ts) =
          Term.list_comb (f $ Syntax_Trans.quote_tr' syntax_const‹_antiquote› t, ts)
      | quote_tr' _ _ = raise Match;

    val assert_tr' = quote_tr' (Syntax.const syntax_const‹_Assert›);

    fun bexp_tr' name ((Const (const_syntaxCollect, _) $ t) :: ts) =
          quote_tr' (Syntax.const name) (t :: ts)
      | bexp_tr' _ _ = raise Match;

    fun assign_tr' (Abs (x, _, f $ k $ Bound 0) :: ts) =
          quote_tr' (Syntax.const syntax_const‹_Assign› $ Syntax_Trans.update_name_tr' f)
            (Abs (x, dummyT, Syntax_Trans.const_abs_tr' k) :: ts)
      | assign_tr' _ = raise Match;
  in
   [(const_syntaxCollect, K assert_tr'),
    (const_syntaxBasic, K assign_tr'),
    (const_syntaxCond, K (bexp_tr' syntax_const‹_Cond›)),
    (const_syntaxWhile, K (bexp_tr' syntax_const‹_While›))]
  end

end