author | wenzelm |
Mon, 03 Feb 2025 20:22:51 +0100 | |
changeset 82073 | 879be333e939 |
parent 80914 | d97fdabd9e2b |
permissions | -rw-r--r-- |
59189 | 1 |
section \<open>Concrete Syntax\<close> |
13020 | 2 |
|
15425 | 3 |
theory RG_Syntax |
4 |
imports RG_Hoare Quote_Antiquote |
|
5 |
begin |
|
13020 | 6 |
|
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
69597
diff
changeset
|
7 |
abbreviation Skip :: "'a com" (\<open>SKIP\<close>) |
35107 | 8 |
where "SKIP \<equiv> Basic id" |
9 |
||
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
69597
diff
changeset
|
10 |
notation Seq (\<open>(_;;/ _)\<close> [60,61] 60) |
35107 | 11 |
|
13020 | 12 |
syntax |
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
69597
diff
changeset
|
13 |
"_Assign" :: "idt \<Rightarrow> 'b \<Rightarrow> 'a com" (\<open>(\<acute>_ :=/ _)\<close> [70, 65] 61) |
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
69597
diff
changeset
|
14 |
"_Cond" :: "'a bexp \<Rightarrow> 'a com \<Rightarrow> 'a com \<Rightarrow> 'a com" (\<open>(0IF _/ THEN _/ ELSE _/FI)\<close> [0, 0, 0] 61) |
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
69597
diff
changeset
|
15 |
"_Cond2" :: "'a bexp \<Rightarrow> 'a com \<Rightarrow> 'a com" (\<open>(0IF _ THEN _ FI)\<close> [0,0] 56) |
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
69597
diff
changeset
|
16 |
"_While" :: "'a bexp \<Rightarrow> 'a com \<Rightarrow> 'a com" (\<open>(0WHILE _ /DO _ /OD)\<close> [0, 0] 61) |
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
69597
diff
changeset
|
17 |
"_Await" :: "'a bexp \<Rightarrow> 'a com \<Rightarrow> 'a com" (\<open>(0AWAIT _ /THEN /_ /END)\<close> [0,0] 61) |
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
69597
diff
changeset
|
18 |
"_Atom" :: "'a com \<Rightarrow> 'a com" (\<open>(\<langle>_\<rangle>)\<close> 61) |
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
69597
diff
changeset
|
19 |
"_Wait" :: "'a bexp \<Rightarrow> 'a com" (\<open>(0WAIT _ END)\<close> 61) |
13020 | 20 |
|
21 |
translations |
|
35107 | 22 |
"\<acute>x := a" \<rightharpoonup> "CONST Basic \<guillemotleft>\<acute>(_update_name x (\<lambda>_. a))\<guillemotright>" |
53241 | 23 |
"IF b THEN c1 ELSE c2 FI" \<rightharpoonup> "CONST Cond \<lbrace>b\<rbrace> c1 c2" |
13020 | 24 |
"IF b THEN c FI" \<rightleftharpoons> "IF b THEN c ELSE SKIP FI" |
53241 | 25 |
"WHILE b DO c OD" \<rightharpoonup> "CONST While \<lbrace>b\<rbrace> c" |
26 |
"AWAIT b THEN c END" \<rightleftharpoons> "CONST Await \<lbrace>b\<rbrace> c" |
|
35113 | 27 |
"\<langle>c\<rangle>" \<rightleftharpoons> "AWAIT CONST True THEN c END" |
13020 | 28 |
"WAIT b END" \<rightleftharpoons> "AWAIT b THEN SKIP END" |
29 |
||
41229
d797baa3d57c
replaced command 'nonterminals' by slightly modernized version 'nonterminal';
wenzelm
parents:
35145
diff
changeset
|
30 |
nonterminal prgs |
13020 | 31 |
|
32 |
syntax |
|
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
69597
diff
changeset
|
33 |
"_PAR" :: "prgs \<Rightarrow> 'a" (\<open>COBEGIN//_//COEND\<close> 60) |
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
69597
diff
changeset
|
34 |
"_prg" :: "'a \<Rightarrow> prgs" (\<open>_\<close> 57) |
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
69597
diff
changeset
|
35 |
"_prgs" :: "['a, prgs] \<Rightarrow> prgs" (\<open>_//\<parallel>//_\<close> [60,57] 57) |
13020 | 36 |
|
37 |
translations |
|
38 |
"_prg a" \<rightharpoonup> "[a]" |
|
39 |
"_prgs a ps" \<rightharpoonup> "a # ps" |
|
40 |
"_PAR ps" \<rightharpoonup> "ps" |
|
41 |
||
42 |
syntax |
|
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
69597
diff
changeset
|
43 |
"_prg_scheme" :: "['a, 'a, 'a, 'a] \<Rightarrow> prgs" (\<open>SCHEME [_ \<le> _ < _] _\<close> [0,0,0,60] 57) |
13020 | 44 |
|
45 |
translations |
|
34940 | 46 |
"_prg_scheme j i k c" \<rightleftharpoons> "(CONST map (\<lambda>i. c) [j..<k])" |
13020 | 47 |
|
59189 | 48 |
text \<open>Translations for variables before and after a transition:\<close> |
13020 | 49 |
|
59189 | 50 |
syntax |
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
69597
diff
changeset
|
51 |
"_before" :: "id \<Rightarrow> 'a" (\<open>\<ordmasculine>_\<close>) |
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
69597
diff
changeset
|
52 |
"_after" :: "id \<Rightarrow> 'a" (\<open>\<ordfeminine>_\<close>) |
59189 | 53 |
|
13020 | 54 |
translations |
35107 | 55 |
"\<ordmasculine>x" \<rightleftharpoons> "x \<acute>CONST fst" |
56 |
"\<ordfeminine>x" \<rightleftharpoons> "x \<acute>CONST snd" |
|
13020 | 57 |
|
59189 | 58 |
print_translation \<open> |
13020 | 59 |
let |
60 |
fun quote_tr' f (t :: ts) = |
|
69597 | 61 |
Term.list_comb (f $ Syntax_Trans.quote_tr' \<^syntax_const>\<open>_antiquote\<close> t, ts) |
13020 | 62 |
| quote_tr' _ _ = raise Match; |
63 |
||
69597 | 64 |
val assert_tr' = quote_tr' (Syntax.const \<^syntax_const>\<open>_Assert\<close>); |
13020 | 65 |
|
69597 | 66 |
fun bexp_tr' name ((Const (\<^const_syntax>\<open>Collect\<close>, _) $ t) :: ts) = |
13020 | 67 |
quote_tr' (Syntax.const name) (t :: ts) |
68 |
| bexp_tr' _ _ = raise Match; |
|
69 |
||
25706 | 70 |
fun assign_tr' (Abs (x, _, f $ k $ Bound 0) :: ts) = |
69597 | 71 |
quote_tr' (Syntax.const \<^syntax_const>\<open>_Assign\<close> $ Syntax_Trans.update_name_tr' f) |
42284 | 72 |
(Abs (x, dummyT, Syntax_Trans.const_abs_tr' k) :: ts) |
13020 | 73 |
| assign_tr' _ = raise Match; |
74 |
in |
|
69597 | 75 |
[(\<^const_syntax>\<open>Collect\<close>, K assert_tr'), |
76 |
(\<^const_syntax>\<open>Basic\<close>, K assign_tr'), |
|
77 |
(\<^const_syntax>\<open>Cond\<close>, K (bexp_tr' \<^syntax_const>\<open>_Cond\<close>)), |
|
78 |
(\<^const_syntax>\<open>While\<close>, K (bexp_tr' \<^syntax_const>\<open>_While\<close>))] |
|
13020 | 79 |
end |
59189 | 80 |
\<close> |
13020 | 81 |
|
62390 | 82 |
end |