Theory picore_SIMP_Syntax

theory picore_SIMP_Syntax
imports picore_SIMP
(*
Created by Yongwang Zhao (zhaoyongwang@gmail.com, zhaoyw@buaa.edu.cn)
School of Computer Science & Engineering, Beihang University, China
*)

section ‹Concrete Syntax of PiCore-SIMP›

theory picore_SIMP_Syntax
imports picore_SIMP (*PiCore_RG_IFS *)

(*keywords "procedures" :: thy_decl*)

begin

syntax
  "_quote"     :: "'b ⇒ ('s ⇒ 'b)"                ("(«_»)" [0] 1000)
  "_antiquote" :: "('s ⇒ 'b) ⇒ 'b"                ("´_" [1000] 1000)
  "_Assert"    :: "'s ⇒ 's set"                    ("(⦃_⦄)" [0] 1000)

translations
  "⦃b⦄"  "CONST Collect «b»"

parse_translation ‹
  let
    fun quote_tr [t] = Syntax_Trans.quote_tr @{syntax_const "_antiquote"} t
      | quote_tr ts = raise TERM ("quote_tr", ts);
  in [(@{syntax_const "_quote"}, K quote_tr)] end
›

definition Skip :: "'s com"  ("SKIP")
  where "SKIP ≡ Basic id"

notation Seq  ("(_;;/ _)" [60,61] 60)

(*
syntax
rghoare_p :: "'Env ⇒ 'prog ⇒ ['s set, ('s × 's) set, ('s × 's) set, 's set] ⇒ bool"
    ("_ ⊢ _ satp [_, _, _, _]" [60,60,0,0,0,0] 45)
rghoare_e :: "'Env ⇒ ('l,'k,'s,'prog) event ⇒ ['s set, ('s × 's) set, ('s × 's) set, 's set] ⇒ bool"
    ("_ ⊢ _ sate [_, _, _, _]" [60,60,0,0,0,0] 45)
Evt_sat_RG:: "'Env ⇒ ('l,'k,'s,'prog) event ⇒ 's rgformula ⇒ bool" ("(_ _⊢_)" [60,60,60] 61)
rghoare_es :: "'Env ⇒ ('l,'k,'s,'prog) rgformula_ess ⇒ ['s set, ('s × 's) set, ('s × 's) set, 's set] ⇒ bool"
    ("_ ⊢ _ sats [_, _, _, _]" [60,60,0,0,0,0] 45)
rghoare_pes :: "'Env ⇒ ('l,'k,'s,'prog) rgformula_par ⇒ ['s set, ('s × 's) set, ('s × 's) set, 's set] ⇒ bool"
          ("_ ⊢ _ SAT [_, _, _, _]" [60,60,0,0,0,0] 45)
Evt_sat_RG:: "'Env ⇒ ('l,'k,'s,'prog) event ⇒ 's rgformula ⇒ bool" ("(_ _⊢_)" [60,60,60] 61)
Esys_sat_RG :: "'Env ⇒ ('l,'k,'s,'prog) rgformula_ess ⇒ 's rgformula ⇒ bool" ("(_ _⊢es_)" [60,60,60] 61)
*)

syntax
  "_Assign"    :: "idt ⇒ 'b ⇒ 's com"                      ("(´_ :=/ _)" [70, 65] 61)
  "_Cond"      :: "'s bexp ⇒ 's com ⇒ 's com ⇒ 's com"  ("(0IF _/ THEN _/ ELSE _/FI)" [0, 0, 0] 61)
  "_Cond2"     :: "'s bexp ⇒ 's com ⇒ 's com"             ("(0IF _ THEN _ FI)" [0,0] 62(*56*))
  "_While"     :: "'s bexp ⇒ 's com ⇒ 's com"             ("(0WHILE _ /DO _ /OD)"  [0, 0] 61)
  "_Await"     :: "'s bexp ⇒ 's com ⇒ 's com"             ("(0AWAIT _ /THEN /_ /END)"  [0,0] 61)
  "_Atom"      :: "'s com ⇒ 's com"                        ("(0ATOMIC _ END)" 61)(*("(⟨ _ ⟩)" 61)*)
  "_Wait"      :: "'s bexp ⇒ 's com"                       ("(0WAIT _ END)" 61)
  "_For"       :: "'s com ⇒ 's bexp ⇒ 's com ⇒ 's com ⇒ 's com" ("(0FOR _;/ _;/ _/ DO _/ ROF)")
  "_Event"     :: "['a, 'a, 'a] ⇒ ('l,'s,'s com option) event" ("(EVENT _ WHEN _ THEN _ END)" [0,0,0] 61)
  "_Event2"     :: "['a, 'a] ⇒ ('l,'s,'s com option) event" ("(EVENT _ THEN _ END)" [0,0] 61)
  "_Event_a"     :: "['a, 'a, 'a] ⇒ ('l,'s,'s com option) event" ("(EVENTA _ WHEN _ THEN _ END)" [0,0,0] 61)
  "_Event_a2"     :: "['a, 'a] ⇒ ('l,'s,'s com option) event" ("(EVENTA _ THEN _ END)" [0,0] 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"*)
  "ATOMIC c END"  "AWAIT CONST True THEN c END"
  "WAIT b END"  "AWAIT b THEN SKIP END"
  "FOR a; b; c DO p ROF"  "a;; WHILE b DO p;;c OD"
  "EVENT l WHEN g THEN bd END"  "CONST EBasic (l, ⦃g⦄, CONST Some bd)"
  "EVENT l THEN bd END"  "EVENT l WHEN CONST True THEN bd END"
  "EVENTA l WHEN g THEN bd END"  "CONST EAtom (l, ⦃g⦄, CONST Some bd)"
  "EVENTA l THEN bd END"  "EVENTA l WHEN CONST True THEN bd END"

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_syntax Collect}, _) $ 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_syntax Collect}, K assert_tr'),
    (@{const_syntax Basic}, K assign_tr'),
    (@{const_syntax Cond}, K (bexp_tr' @{syntax_const "_Cond"})),
    (@{const_syntax While}, K (bexp_tr' @{syntax_const "_While"}))]
  end
›

lemma colltrue_eq_univ[simp]: "⦃True⦄ = UNIV" by auto

end