products/sources/formale sprachen/Isabelle/HOL/Hoare_Parallel image not shown  

Quellcode-Bibliothek

© Kompilation durch diese Firma

[Weder Korrektheit noch Funktionsfähigkeit der Software werden zugesichert.]

Datei: RG_Syntax.thy   Sprache: Isabelle

Original von: Isabelle©

section \<open>Concrete Syntax\<close>

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" \<rightharpoonup> "CONST Cond \<lbrace>b\<rbrace> c1 c2"
  "IF b THEN c FI" \<rightleftharpoons> "IF b THEN c ELSE SKIP FI"
  "WHILE b DO c OD" \<rightharpoonup> "CONST While \<lbrace>b\<rbrace> c"
  "AWAIT b THEN c END" \<rightleftharpoons> "CONST Await \<lbrace>b\<rbrace> c"
  "\c\" \ "AWAIT CONST True THEN c END"
  "WAIT b END" \<rightleftharpoons> "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" \<rightharpoonup> "[a]"
  "_prgs a ps" \<rightharpoonup> "a # ps"
  "_PAR ps" \<rightharpoonup> "ps"

syntax
  "_prg_scheme" :: "['a, 'a, 'a, 'a] \ prgs" ("SCHEME [_ \ _ < _] _" [0,0,0,60] 57)

translations
  "_prg_scheme j i k c" \<rightleftharpoons> "(CONST map (\<lambda>i. c) [j..<k])"

text \<open>Translations for variables before and after a transition:\<close>

syntax
  "_before" :: "id \ 'a" ("\_")
  "_after"  :: "id \ 'a" ("\_")

translations
  "\x" \ "x \CONST fst"
  "\x" \ "x \CONST snd"

print_translation \<open>
  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>\<open>_Assert\<close>);

    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>\<open>Collect\<close>, K assert_tr'),
    (\<^const_syntax>\<open>Basic\<close>, K assign_tr'),
    (\<^const_syntax>\<open>Cond\<close>, K (bexp_tr' \<^syntax_const>\<open>_Cond\<close>)),
    (\<^const_syntax>\<open>While\<close>, K (bexp_tr' \<^syntax_const>\<open>_While\<close>))]
  end
\<close>

end

¤ Dauer der Verarbeitung: 0.1 Sekunden  (vorverarbeitet)  ¤





Download des
Quellennavigators
Download des
sprechenden Kalenders

in der Quellcodebibliothek suchen




Haftungshinweis

Die Informationen auf dieser Webseite wurden nach bestem Wissen sorgfältig zusammengestellt. Es wird jedoch weder Vollständigkeit, noch Richtigkeit, noch Qualität der bereit gestellten Informationen zugesichert.


Bemerkung:

Die farbliche Syntaxdarstellung ist noch experimentell.


Bot Zugriff