src/HOL/Hoare/hoare_syntax.ML
author wenzelm
Wed, 22 Apr 2020 19:22:43 +0200
changeset 71787 acfe72ff00c2
parent 69597 ff784d5a5bfb
child 72806 4fa08e083865
permissions -rw-r--r--
merged
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
42153
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
     1
(*  Title:      HOL/Hoare/hoare_syntax.ML
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
     2
    Author:     Leonor Prensa Nieto & Tobias Nipkow
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
     3
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
     4
Syntax translations for Hoare logic.
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
     5
*)
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
     6
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
     7
signature HOARE_SYNTAX =
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
     8
sig
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
     9
  val hoare_vars_tr: term list -> term
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    10
  val spec_tr': string -> term list -> term
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    11
end;
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    12
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    13
structure Hoare_Syntax: HOARE_SYNTAX =
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    14
struct
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    15
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    16
(** parse translation **)
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    17
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    18
local
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    19
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    20
fun idt_name (Free (x, _)) = SOME x
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 61424
diff changeset
    21
  | idt_name (Const (\<^syntax_const>\<open>_constrain\<close>, _) $ t $ _) = idt_name t
42153
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    22
  | idt_name _ = NONE;
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    23
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    24
fun eq_idt tu =
59058
a78612c67ec0 renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents: 55659
diff changeset
    25
  (case apply2 idt_name tu of
42153
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    26
    (SOME x, SOME y) => x = y
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    27
  | _ => false);
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    28
42284
326f57825e1a explicit structure Syntax_Trans;
wenzelm
parents: 42153
diff changeset
    29
fun mk_abstuple [x] body = Syntax_Trans.abs_tr [x, body]
42153
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    30
  | mk_abstuple (x :: xs) body =
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 61424
diff changeset
    31
      Syntax.const \<^const_syntax>\<open>case_prod\<close> $ Syntax_Trans.abs_tr [x, mk_abstuple xs body];
42153
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    32
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    33
fun mk_fbody x e [y] = if eq_idt (x, y) then e else y
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    34
  | mk_fbody x e (y :: xs) =
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 61424
diff changeset
    35
      Syntax.const \<^const_syntax>\<open>Pair\<close> $
42153
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    36
        (if eq_idt (x, y) then e else y) $ mk_fbody x e xs;
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    37
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    38
fun mk_fexp x e xs = mk_abstuple xs (mk_fbody x e xs);
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    39
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    40
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    41
(* bexp_tr & assn_tr *)
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    42
(*all meta-variables for bexp except for TRUE are translated as if they
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    43
  were boolean expressions*)
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    44
55659
4089f6e98ab9 reduced ML warnings;
wenzelm
parents: 55414
diff changeset
    45
fun bexp_tr (Const ("TRUE", _)) _ = Syntax.const "TRUE"   (* FIXME !? *)
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 61424
diff changeset
    46
  | bexp_tr b xs = Syntax.const \<^const_syntax>\<open>Collect\<close> $ mk_abstuple xs b;
42153
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    47
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 61424
diff changeset
    48
fun assn_tr r xs = Syntax.const \<^const_syntax>\<open>Collect\<close> $ mk_abstuple xs r;
42153
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    49
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    50
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    51
(* com_tr *)
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    52
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 61424
diff changeset
    53
fun com_tr (Const (\<^syntax_const>\<open>_assign\<close>, _) $ x $ e) xs =
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 61424
diff changeset
    54
      Syntax.const \<^const_syntax>\<open>Basic\<close> $ mk_fexp x e xs
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 61424
diff changeset
    55
  | com_tr (Const (\<^const_syntax>\<open>Basic\<close>,_) $ f) _ = Syntax.const \<^const_syntax>\<open>Basic\<close> $ f
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 61424
diff changeset
    56
  | com_tr (Const (\<^const_syntax>\<open>Seq\<close>,_) $ c1 $ c2) xs =
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 61424
diff changeset
    57
      Syntax.const \<^const_syntax>\<open>Seq\<close> $ com_tr c1 xs $ com_tr c2 xs
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 61424
diff changeset
    58
  | com_tr (Const (\<^const_syntax>\<open>Cond\<close>,_) $ b $ c1 $ c2) xs =
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 61424
diff changeset
    59
      Syntax.const \<^const_syntax>\<open>Cond\<close> $ bexp_tr b xs $ com_tr c1 xs $ com_tr c2 xs
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 61424
diff changeset
    60
  | com_tr (Const (\<^const_syntax>\<open>While\<close>,_) $ b $ I $ c) xs =
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 61424
diff changeset
    61
      Syntax.const \<^const_syntax>\<open>While\<close> $ bexp_tr b xs $ assn_tr I xs $ com_tr c xs
42153
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    62
  | com_tr t _ = t;
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    63
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 61424
diff changeset
    64
fun vars_tr (Const (\<^syntax_const>\<open>_idts\<close>, _) $ idt $ vars) = idt :: vars_tr vars
42153
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    65
  | vars_tr t = [t];
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    66
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    67
in
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    68
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    69
fun hoare_vars_tr [vars, pre, prg, post] =
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    70
      let val xs = vars_tr vars
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 61424
diff changeset
    71
      in Syntax.const \<^const_syntax>\<open>Valid\<close> $
42153
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    72
         assn_tr pre xs $ com_tr prg xs $ assn_tr post xs
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    73
      end
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    74
  | hoare_vars_tr ts = raise TERM ("hoare_vars_tr", ts);
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    75
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    76
end;
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    77
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    78
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    79
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    80
(** print translation **)
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    81
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    82
local
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    83
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    84
fun dest_abstuple
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 61424
diff changeset
    85
      (Const (\<^const_syntax>\<open>case_prod\<close>, _) $ Abs (v, _, body)) =
42153
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    86
        subst_bound (Syntax.free v, dest_abstuple body)
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    87
  | dest_abstuple (Abs (v,_, body)) = subst_bound (Syntax.free v, body)
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    88
  | dest_abstuple tm = tm;
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    89
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 61424
diff changeset
    90
fun abs2list (Const (\<^const_syntax>\<open>case_prod\<close>, _) $ Abs (x, T, t)) = Free (x, T) :: abs2list t
55659
4089f6e98ab9 reduced ML warnings;
wenzelm
parents: 55414
diff changeset
    91
  | abs2list (Abs (x, T, _)) = [Free (x, T)]
42153
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    92
  | abs2list _ = [];
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    93
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 61424
diff changeset
    94
fun mk_ts (Const (\<^const_syntax>\<open>case_prod\<close>, _) $ Abs (_, _, t)) = mk_ts t
55659
4089f6e98ab9 reduced ML warnings;
wenzelm
parents: 55414
diff changeset
    95
  | mk_ts (Abs (_, _, t)) = mk_ts t
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 61424
diff changeset
    96
  | mk_ts (Const (\<^const_syntax>\<open>Pair\<close>, _) $ a $ b) = a :: mk_ts b
42153
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    97
  | mk_ts t = [t];
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    98
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 61424
diff changeset
    99
fun mk_vts (Const (\<^const_syntax>\<open>case_prod\<close>,_) $ Abs (x, _, t)) =
42153
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
   100
      (Syntax.free x :: abs2list t, mk_ts t)
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
   101
  | mk_vts (Abs (x, _, t)) = ([Syntax.free x], [t])
55659
4089f6e98ab9 reduced ML warnings;
wenzelm
parents: 55414
diff changeset
   102
  | mk_vts _ = raise Match;
42153
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
   103
55659
4089f6e98ab9 reduced ML warnings;
wenzelm
parents: 55414
diff changeset
   104
fun find_ch [] _ _ = (false, (Syntax.free "not_ch", Syntax.free "not_ch"))  (* FIXME no_ch!? *)
42153
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
   105
  | find_ch ((v, t) :: vts) i xs =
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
   106
      if t = Bound i then find_ch vts (i - 1) xs
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
   107
      else (true, (v, subst_bounds (xs, t)));
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
   108
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 61424
diff changeset
   109
fun is_f (Const (\<^const_syntax>\<open>case_prod\<close>, _) $ Abs _) = true
55659
4089f6e98ab9 reduced ML warnings;
wenzelm
parents: 55414
diff changeset
   110
  | is_f (Abs _) = true
4089f6e98ab9 reduced ML warnings;
wenzelm
parents: 55414
diff changeset
   111
  | is_f _ = false;
42153
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
   112
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
   113
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
   114
(* assn_tr' & bexp_tr'*)
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
   115
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 61424
diff changeset
   116
fun assn_tr' (Const (\<^const_syntax>\<open>Collect\<close>, _) $ T) = dest_abstuple T
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 61424
diff changeset
   117
  | assn_tr' (Const (\<^const_syntax>\<open>inter\<close>, _) $
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 61424
diff changeset
   118
        (Const (\<^const_syntax>\<open>Collect\<close>, _) $ T1) $ (Const (\<^const_syntax>\<open>Collect\<close>, _) $ T2)) =
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 61424
diff changeset
   119
      Syntax.const \<^const_syntax>\<open>inter\<close> $ dest_abstuple T1 $ dest_abstuple T2
42153
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
   120
  | assn_tr' t = t;
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
   121
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 61424
diff changeset
   122
fun bexp_tr' (Const (\<^const_syntax>\<open>Collect\<close>, _) $ T) = dest_abstuple T
42153
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
   123
  | bexp_tr' t = t;
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
   124
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
   125
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
   126
(* com_tr' *)
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
   127
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
   128
fun mk_assign f =
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
   129
  let
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
   130
    val (vs, ts) = mk_vts f;
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
   131
    val (ch, which) = find_ch (vs ~~ ts) (length vs - 1) (rev vs);
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
   132
  in
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
   133
    if ch
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 61424
diff changeset
   134
    then Syntax.const \<^syntax_const>\<open>_assign\<close> $ fst which $ snd which
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 61424
diff changeset
   135
    else Syntax.const \<^const_syntax>\<open>annskip\<close>
42153
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
   136
  end;
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
   137
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 61424
diff changeset
   138
fun com_tr' (Const (\<^const_syntax>\<open>Basic\<close>, _) $ f) =
42153
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
   139
      if is_f f then mk_assign f
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 61424
diff changeset
   140
      else Syntax.const \<^const_syntax>\<open>Basic\<close> $ f
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 61424
diff changeset
   141
  | com_tr' (Const (\<^const_syntax>\<open>Seq\<close>,_) $ c1 $ c2) =
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 61424
diff changeset
   142
      Syntax.const \<^const_syntax>\<open>Seq\<close> $ com_tr' c1 $ com_tr' c2
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 61424
diff changeset
   143
  | com_tr' (Const (\<^const_syntax>\<open>Cond\<close>, _) $ b $ c1 $ c2) =
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 61424
diff changeset
   144
      Syntax.const \<^const_syntax>\<open>Cond\<close> $ bexp_tr' b $ com_tr' c1 $ com_tr' c2
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 61424
diff changeset
   145
  | com_tr' (Const (\<^const_syntax>\<open>While\<close>, _) $ b $ I $ c) =
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 61424
diff changeset
   146
      Syntax.const \<^const_syntax>\<open>While\<close> $ bexp_tr' b $ assn_tr' I $ com_tr' c
42153
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
   147
  | com_tr' t = t;
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
   148
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
   149
in
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
   150
55659
4089f6e98ab9 reduced ML warnings;
wenzelm
parents: 55414
diff changeset
   151
fun spec_tr' syn [p, c, q] = Syntax.const syn $ assn_tr' p $ com_tr' c $ assn_tr' q;
42153
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
   152
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
   153
end;
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
   154
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
   155
end;
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
   156