src/HOL/Hoare_Parallel/RG_Syntax.thy
changeset 35145 f132a4fd8679
parent 35113 1a0c129bb2e0
child 41229 d797baa3d57c
equal deleted inserted replaced
35144:8b8302da3a55 35145:f132a4fd8679
    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 
    71     fun upd_tr' (x_upd, T) =
       
    72       (case try (unsuffix Record.updateN) x_upd of
       
    73         SOME x => (x, if T = dummyT then T else Term.domain_type T)
       
    74       | NONE => raise Match);
       
    75 
       
    76     fun update_name_tr' (Free x) = Free (upd_tr' x)
       
    77       | update_name_tr' ((c as Const (@{syntax_const "_free"}, _)) $ Free x) =
       
    78           c $ Free (upd_tr' x)
       
    79       | update_name_tr' (Const x) = Const (upd_tr' x)
       
    80       | update_name_tr' _ = raise Match;
       
    81 
       
    82     fun K_tr' (Abs (_, _, t)) =
    71     fun K_tr' (Abs (_, _, t)) =
    83           if null (loose_bnos t) then t else raise Match
    72           if null (loose_bnos t) then t else raise Match
    84       | K_tr' (Abs (_, _, Abs (_, _, t) $ Bound 0)) =
    73       | K_tr' (Abs (_, _, Abs (_, _, t) $ Bound 0)) =
    85           if null (loose_bnos t) then t else raise Match
    74           if null (loose_bnos t) then t else raise Match
    86       | K_tr' _ = raise Match;
    75       | K_tr' _ = raise Match;
    87 
    76 
    88     fun assign_tr' (Abs (x, _, f $ k $ Bound 0) :: ts) =
    77     fun assign_tr' (Abs (x, _, f $ k $ Bound 0) :: ts) =
    89           quote_tr' (Syntax.const @{syntax_const "_Assign"} $ update_name_tr' f)
    78           quote_tr' (Syntax.const @{syntax_const "_Assign"} $ Syntax.update_name_tr' f)
    90             (Abs (x, dummyT, K_tr' k) :: ts)
    79             (Abs (x, dummyT, K_tr' k) :: ts)
    91       | assign_tr' _ = raise Match;
    80       | assign_tr' _ = raise Match;
    92   in
    81   in
    93    [(@{const_syntax Collect}, assert_tr'),
    82    [(@{const_syntax Collect}, assert_tr'),
    94     (@{const_syntax Basic}, assign_tr'),
    83     (@{const_syntax Basic}, assign_tr'),