src/HOL/Hoare_Parallel/RG_Syntax.thy
changeset 35113 1a0c129bb2e0
parent 35107 bdca9f765ee4
child 35145 f132a4fd8679
equal deleted inserted replaced
35112:ff6f60e6ab85 35113:1a0c129bb2e0
    18   "_Atom"      :: "'a com \<Rightarrow> 'a com"                        ("(\<langle>_\<rangle>)" 61)
    18   "_Atom"      :: "'a com \<Rightarrow> 'a com"                        ("(\<langle>_\<rangle>)" 61)
    19   "_Wait"      :: "'a bexp \<Rightarrow> 'a com"                       ("(0WAIT _ END)" 61)
    19   "_Wait"      :: "'a bexp \<Rightarrow> 'a com"                       ("(0WAIT _ END)" 61)
    20 
    20 
    21 translations
    21 translations
    22   "\<acute>x := a" \<rightharpoonup> "CONST Basic \<guillemotleft>\<acute>(_update_name x (\<lambda>_. a))\<guillemotright>"
    22   "\<acute>x := a" \<rightharpoonup> "CONST Basic \<guillemotleft>\<acute>(_update_name x (\<lambda>_. a))\<guillemotright>"
    23   "IF b THEN c1 ELSE c2 FI" \<rightharpoonup> "Cond .{b}. c1 c2"
    23   "IF b THEN c1 ELSE c2 FI" \<rightharpoonup> "CONST Cond .{b}. c1 c2"
    24   "IF b THEN c FI" \<rightleftharpoons> "IF b THEN c ELSE SKIP FI"
    24   "IF b THEN c FI" \<rightleftharpoons> "IF b THEN c ELSE SKIP FI"
    25   "WHILE b DO c OD" \<rightharpoonup> "While .{b}. c"
    25   "WHILE b DO c OD" \<rightharpoonup> "CONST While .{b}. c"
    26   "AWAIT b THEN c END" \<rightleftharpoons> "Await .{b}. c"
    26   "AWAIT b THEN c END" \<rightleftharpoons> "CONST Await .{b}. c"
    27   "\<langle>c\<rangle>" \<rightleftharpoons> "AWAIT True THEN c END"
    27   "\<langle>c\<rangle>" \<rightleftharpoons> "AWAIT CONST True THEN c END"
    28   "WAIT b END" \<rightleftharpoons> "AWAIT b THEN SKIP END"
    28   "WAIT b END" \<rightleftharpoons> "AWAIT b THEN SKIP END"
    29 
    29 
    30 nonterminals
    30 nonterminals
    31   prgs
    31   prgs
    32 
    32 
    57   "\<ordfeminine>x" \<rightleftharpoons> "x \<acute>CONST snd"
    57   "\<ordfeminine>x" \<rightleftharpoons> "x \<acute>CONST snd"
    58 
    58 
    59 print_translation {*
    59 print_translation {*
    60   let
    60   let
    61     fun quote_tr' f (t :: ts) =
    61     fun quote_tr' f (t :: ts) =
    62           Term.list_comb (f $ Syntax.quote_tr' "_antiquote" t, ts)
    62           Term.list_comb (f $ Syntax.quote_tr' @{syntax_const "_antiquote"} t, ts)
    63       | quote_tr' _ _ = raise Match;
    63       | quote_tr' _ _ = raise Match;
    64 
    64 
    65     val assert_tr' = quote_tr' (Syntax.const "_Assert");
    65     val assert_tr' = quote_tr' (Syntax.const @{syntax_const "_Assert"});
    66 
    66 
    67     fun bexp_tr' name ((Const (@{const_syntax Collect}, _) $ t) :: ts) =
    67     fun bexp_tr' name ((Const (@{const_syntax Collect}, _) $ t) :: ts) =
    68           quote_tr' (Syntax.const name) (t :: ts)
    68           quote_tr' (Syntax.const name) (t :: ts)
    69       | bexp_tr' _ _ = raise Match;
    69       | bexp_tr' _ _ = raise Match;
    70 
    70 
    72       (case try (unsuffix Record.updateN) x_upd of
    72       (case try (unsuffix Record.updateN) x_upd of
    73         SOME x => (x, if T = dummyT then T else Term.domain_type T)
    73         SOME x => (x, if T = dummyT then T else Term.domain_type T)
    74       | NONE => raise Match);
    74       | NONE => raise Match);
    75 
    75 
    76     fun update_name_tr' (Free x) = Free (upd_tr' x)
    76     fun update_name_tr' (Free x) = Free (upd_tr' x)
    77       | update_name_tr' ((c as Const ("_free", _)) $ Free x) =
    77       | update_name_tr' ((c as Const (@{syntax_const "_free"}, _)) $ Free x) =
    78           c $ Free (upd_tr' x)
    78           c $ Free (upd_tr' x)
    79       | update_name_tr' (Const x) = Const (upd_tr' x)
    79       | update_name_tr' (Const x) = Const (upd_tr' x)
    80       | update_name_tr' _ = raise Match;
    80       | update_name_tr' _ = raise Match;
    81 
    81 
    82     fun K_tr' (Abs (_, _, t)) =
    82     fun K_tr' (Abs (_, _, t)) =
    84       | K_tr' (Abs (_, _, Abs (_, _, t) $ Bound 0)) =
    84       | K_tr' (Abs (_, _, Abs (_, _, t) $ Bound 0)) =
    85           if null (loose_bnos t) then t else raise Match
    85           if null (loose_bnos t) then t else raise Match
    86       | K_tr' _ = raise Match;
    86       | K_tr' _ = raise Match;
    87 
    87 
    88     fun assign_tr' (Abs (x, _, f $ k $ Bound 0) :: ts) =
    88     fun assign_tr' (Abs (x, _, f $ k $ Bound 0) :: ts) =
    89           quote_tr' (Syntax.const "_Assign" $ update_name_tr' f)
    89           quote_tr' (Syntax.const @{syntax_const "_Assign"} $ update_name_tr' f)
    90             (Abs (x, dummyT, K_tr' k) :: ts)
    90             (Abs (x, dummyT, K_tr' k) :: ts)
    91       | assign_tr' _ = raise Match;
    91       | assign_tr' _ = raise Match;
    92   in
    92   in
    93    [(@{const_syntax Collect}, assert_tr'),
    93    [(@{const_syntax Collect}, assert_tr'),
    94     (@{const_syntax Basic}, assign_tr'),
    94     (@{const_syntax Basic}, assign_tr'),
    95     (@{const_syntax Cond}, bexp_tr' "_Cond"),
    95     (@{const_syntax Cond}, bexp_tr' @{syntax_const "_Cond"}),
    96     (@{const_syntax While}, bexp_tr' "_While_inv")]
    96     (@{const_syntax While}, bexp_tr' @{syntax_const "_While"})]
    97   end
    97   end
    98 *}
    98 *}
    99 
    99 
   100 end
   100 end