src/HOL/Hoare/hoare_syntax.ML
author nipkow
Tue, 12 Oct 2021 20:57:43 +0200
changeset 74503 403ce50e6a2a
parent 72998 7ea253f93606
child 74561 8e6c973003c8
permissions -rw-r--r--
separated commands from annotations to be able to abstract about the latter only
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
74503
403ce50e6a2a separated commands from annotations to be able to abstract about the latter only
nipkow
parents: 72998
diff changeset
     3
    Author:     Walter Guttmann (initial extension to total-correctness proofs)
403ce50e6a2a separated commands from annotations to be able to abstract about the latter only
nipkow
parents: 72998
diff changeset
     4
    Author:     Tobias Nipkow (complete version of total correctness with separate anno type)
42153
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
Syntax translations for Hoare logic.
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
     7
*)
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
     8
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
     9
signature HOARE_SYNTAX =
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    10
sig
72987
b1be35908165 clarified modules: avoid multiple uses of the same ML file;
wenzelm
parents: 72806
diff changeset
    11
  val hoare_vars_tr: Proof.context -> term list -> term
b1be35908165 clarified modules: avoid multiple uses of the same ML file;
wenzelm
parents: 72806
diff changeset
    12
  val hoare_tc_vars_tr: Proof.context -> term list -> term
b1be35908165 clarified modules: avoid multiple uses of the same ML file;
wenzelm
parents: 72806
diff changeset
    13
  val spec_tr': string -> Proof.context -> term list -> term
b1be35908165 clarified modules: avoid multiple uses of the same ML file;
wenzelm
parents: 72806
diff changeset
    14
  val setup:
b1be35908165 clarified modules: avoid multiple uses of the same ML file;
wenzelm
parents: 72806
diff changeset
    15
    {Basic: string, Skip: string, Seq: string, Cond: string, While: string,
b1be35908165 clarified modules: avoid multiple uses of the same ML file;
wenzelm
parents: 72806
diff changeset
    16
      Valid: string, ValidTC: string} -> theory -> theory
42153
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    17
end;
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    18
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    19
structure Hoare_Syntax: HOARE_SYNTAX =
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    20
struct
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    21
72987
b1be35908165 clarified modules: avoid multiple uses of the same ML file;
wenzelm
parents: 72806
diff changeset
    22
(** theory data **)
b1be35908165 clarified modules: avoid multiple uses of the same ML file;
wenzelm
parents: 72806
diff changeset
    23
b1be35908165 clarified modules: avoid multiple uses of the same ML file;
wenzelm
parents: 72806
diff changeset
    24
structure Data = Theory_Data
b1be35908165 clarified modules: avoid multiple uses of the same ML file;
wenzelm
parents: 72806
diff changeset
    25
(
b1be35908165 clarified modules: avoid multiple uses of the same ML file;
wenzelm
parents: 72806
diff changeset
    26
  type T =
b1be35908165 clarified modules: avoid multiple uses of the same ML file;
wenzelm
parents: 72806
diff changeset
    27
    {Basic: string, Skip: string, Seq: string, Cond: string, While: string,
b1be35908165 clarified modules: avoid multiple uses of the same ML file;
wenzelm
parents: 72806
diff changeset
    28
      Valid: string, ValidTC: string} option;
b1be35908165 clarified modules: avoid multiple uses of the same ML file;
wenzelm
parents: 72806
diff changeset
    29
  val empty: T = NONE;
b1be35908165 clarified modules: avoid multiple uses of the same ML file;
wenzelm
parents: 72806
diff changeset
    30
  val extend = I;
b1be35908165 clarified modules: avoid multiple uses of the same ML file;
wenzelm
parents: 72806
diff changeset
    31
  fun merge (data1, data2) =
b1be35908165 clarified modules: avoid multiple uses of the same ML file;
wenzelm
parents: 72806
diff changeset
    32
    if is_some data1 andalso is_some data2 andalso data1 <> data2 then
b1be35908165 clarified modules: avoid multiple uses of the same ML file;
wenzelm
parents: 72806
diff changeset
    33
      error "Cannot merge syntax from different Hoare Logics"
b1be35908165 clarified modules: avoid multiple uses of the same ML file;
wenzelm
parents: 72806
diff changeset
    34
    else merge_options (data1, data2);
b1be35908165 clarified modules: avoid multiple uses of the same ML file;
wenzelm
parents: 72806
diff changeset
    35
);
b1be35908165 clarified modules: avoid multiple uses of the same ML file;
wenzelm
parents: 72806
diff changeset
    36
b1be35908165 clarified modules: avoid multiple uses of the same ML file;
wenzelm
parents: 72806
diff changeset
    37
fun const_syntax ctxt which =
b1be35908165 clarified modules: avoid multiple uses of the same ML file;
wenzelm
parents: 72806
diff changeset
    38
  (case Data.get (Proof_Context.theory_of ctxt) of
b1be35908165 clarified modules: avoid multiple uses of the same ML file;
wenzelm
parents: 72806
diff changeset
    39
    SOME consts => which consts
b1be35908165 clarified modules: avoid multiple uses of the same ML file;
wenzelm
parents: 72806
diff changeset
    40
  | NONE => error "Undefined Hoare syntax consts");
b1be35908165 clarified modules: avoid multiple uses of the same ML file;
wenzelm
parents: 72806
diff changeset
    41
b1be35908165 clarified modules: avoid multiple uses of the same ML file;
wenzelm
parents: 72806
diff changeset
    42
val syntax_const = Syntax.const oo const_syntax;
b1be35908165 clarified modules: avoid multiple uses of the same ML file;
wenzelm
parents: 72806
diff changeset
    43
b1be35908165 clarified modules: avoid multiple uses of the same ML file;
wenzelm
parents: 72806
diff changeset
    44
b1be35908165 clarified modules: avoid multiple uses of the same ML file;
wenzelm
parents: 72806
diff changeset
    45
42153
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    46
(** parse translation **)
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    47
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    48
local
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
fun idt_name (Free (x, _)) = SOME x
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 61424
diff changeset
    51
  | 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
    52
  | idt_name _ = NONE;
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    53
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    54
fun eq_idt tu =
59058
a78612c67ec0 renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents: 55659
diff changeset
    55
  (case apply2 idt_name tu of
42153
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    56
    (SOME x, SOME y) => x = y
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    57
  | _ => false);
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    58
42284
326f57825e1a explicit structure Syntax_Trans;
wenzelm
parents: 42153
diff changeset
    59
fun mk_abstuple [x] body = Syntax_Trans.abs_tr [x, body]
42153
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    60
  | mk_abstuple (x :: xs) body =
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 61424
diff changeset
    61
      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
    62
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    63
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
    64
  | mk_fbody x e (y :: xs) =
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 61424
diff changeset
    65
      Syntax.const \<^const_syntax>\<open>Pair\<close> $
42153
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    66
        (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
    67
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    68
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
    69
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    70
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    71
(* bexp_tr & assn_tr *)
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    72
(*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
    73
  were boolean expressions*)
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    74
55659
4089f6e98ab9 reduced ML warnings;
wenzelm
parents: 55414
diff changeset
    75
fun bexp_tr (Const ("TRUE", _)) _ = Syntax.const "TRUE"   (* FIXME !? *)
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 61424
diff changeset
    76
  | 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
    77
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 61424
diff changeset
    78
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
    79
72806
4fa08e083865 Extension of session HOL/Hoare with total correctness proof system by Walter Guttmann
nipkow
parents: 69597
diff changeset
    80
fun var_tr v xs = mk_abstuple xs v;
4fa08e083865 Extension of session HOL/Hoare with total correctness proof system by Walter Guttmann
nipkow
parents: 69597
diff changeset
    81
42153
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    82
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    83
(* com_tr *)
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
    84
72987
b1be35908165 clarified modules: avoid multiple uses of the same ML file;
wenzelm
parents: 72806
diff changeset
    85
fun com_tr ctxt =
b1be35908165 clarified modules: avoid multiple uses of the same ML file;
wenzelm
parents: 72806
diff changeset
    86
  let
b1be35908165 clarified modules: avoid multiple uses of the same ML file;
wenzelm
parents: 72806
diff changeset
    87
    fun tr (Const (\<^syntax_const>\<open>_assign\<close>, _) $ x $ e) xs =
74503
403ce50e6a2a separated commands from annotations to be able to abstract about the latter only
nipkow
parents: 72998
diff changeset
    88
          (syntax_const ctxt #Basic $ mk_fexp x e xs, Syntax.const \<^const_syntax>\<open>Abasic\<close>)
72987
b1be35908165 clarified modules: avoid multiple uses of the same ML file;
wenzelm
parents: 72806
diff changeset
    89
      | tr (Const (\<^syntax_const>\<open>_Seq\<close>,_) $ c1 $ c2) xs =
74503
403ce50e6a2a separated commands from annotations to be able to abstract about the latter only
nipkow
parents: 72998
diff changeset
    90
          let val (c1',a1) = tr c1 xs;
403ce50e6a2a separated commands from annotations to be able to abstract about the latter only
nipkow
parents: 72998
diff changeset
    91
              val (c2',a2) = tr c2 xs
403ce50e6a2a separated commands from annotations to be able to abstract about the latter only
nipkow
parents: 72998
diff changeset
    92
          in (syntax_const ctxt #Seq $ c1' $ c2', Syntax.const \<^const_syntax>\<open>Aseq\<close> $ a1 $ a2) end
72987
b1be35908165 clarified modules: avoid multiple uses of the same ML file;
wenzelm
parents: 72806
diff changeset
    93
      | tr (Const (\<^syntax_const>\<open>_Cond\<close>,_) $ b $ c1 $ c2) xs =
74503
403ce50e6a2a separated commands from annotations to be able to abstract about the latter only
nipkow
parents: 72998
diff changeset
    94
          let val (c1',a1) = tr c1 xs;
403ce50e6a2a separated commands from annotations to be able to abstract about the latter only
nipkow
parents: 72998
diff changeset
    95
              val (c2',a2) = tr c2 xs
403ce50e6a2a separated commands from annotations to be able to abstract about the latter only
nipkow
parents: 72998
diff changeset
    96
          in (syntax_const ctxt #Cond $ bexp_tr b xs $ c1' $ c2',
403ce50e6a2a separated commands from annotations to be able to abstract about the latter only
nipkow
parents: 72998
diff changeset
    97
              Syntax.const \<^const_syntax>\<open>Acond\<close> $ a1 $ a2)
403ce50e6a2a separated commands from annotations to be able to abstract about the latter only
nipkow
parents: 72998
diff changeset
    98
          end
403ce50e6a2a separated commands from annotations to be able to abstract about the latter only
nipkow
parents: 72998
diff changeset
    99
      | tr (Const (\<^syntax_const>\<open>_While\<close>,_) $ b $ i $ v $ c) xs =
403ce50e6a2a separated commands from annotations to be able to abstract about the latter only
nipkow
parents: 72998
diff changeset
   100
          let val (c',a) = tr c xs
403ce50e6a2a separated commands from annotations to be able to abstract about the latter only
nipkow
parents: 72998
diff changeset
   101
              val (v',A) = case v of
403ce50e6a2a separated commands from annotations to be able to abstract about the latter only
nipkow
parents: 72998
diff changeset
   102
                Const (\<^const_syntax>\<open>HOL.eq\<close>, _) $ z $ t => (t, Syntax_Trans.abs_tr [z, a]) |
403ce50e6a2a separated commands from annotations to be able to abstract about the latter only
nipkow
parents: 72998
diff changeset
   103
                t => (t, absdummy dummyT a)
403ce50e6a2a separated commands from annotations to be able to abstract about the latter only
nipkow
parents: 72998
diff changeset
   104
          in (syntax_const ctxt #While $ bexp_tr b xs $ c',
403ce50e6a2a separated commands from annotations to be able to abstract about the latter only
nipkow
parents: 72998
diff changeset
   105
              Syntax.const \<^const_syntax>\<open>Awhile\<close>
403ce50e6a2a separated commands from annotations to be able to abstract about the latter only
nipkow
parents: 72998
diff changeset
   106
                $ assn_tr i xs $ var_tr v' xs $ A)
403ce50e6a2a separated commands from annotations to be able to abstract about the latter only
nipkow
parents: 72998
diff changeset
   107
          end
403ce50e6a2a separated commands from annotations to be able to abstract about the latter only
nipkow
parents: 72998
diff changeset
   108
      | tr (Const (\<^syntax_const>\<open>_While0\<close>,_) $ b $ I $ c) xs =
403ce50e6a2a separated commands from annotations to be able to abstract about the latter only
nipkow
parents: 72998
diff changeset
   109
          let val (c',a) = tr c xs
403ce50e6a2a separated commands from annotations to be able to abstract about the latter only
nipkow
parents: 72998
diff changeset
   110
          in (syntax_const ctxt #While $ bexp_tr b xs $ c',
403ce50e6a2a separated commands from annotations to be able to abstract about the latter only
nipkow
parents: 72998
diff changeset
   111
              Syntax.const \<^const_syntax>\<open>Awhile\<close>
403ce50e6a2a separated commands from annotations to be able to abstract about the latter only
nipkow
parents: 72998
diff changeset
   112
                $ assn_tr I xs $ var_tr (Syntax.const \<^const_syntax>\<open>zero_class.zero\<close>) xs
403ce50e6a2a separated commands from annotations to be able to abstract about the latter only
nipkow
parents: 72998
diff changeset
   113
                $ absdummy dummyT a)
403ce50e6a2a separated commands from annotations to be able to abstract about the latter only
nipkow
parents: 72998
diff changeset
   114
          end
403ce50e6a2a separated commands from annotations to be able to abstract about the latter only
nipkow
parents: 72998
diff changeset
   115
      | tr t _ = (t, Syntax.const \<^const_syntax>\<open>Abasic\<close>)
72987
b1be35908165 clarified modules: avoid multiple uses of the same ML file;
wenzelm
parents: 72806
diff changeset
   116
  in tr end;
42153
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
   117
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 61424
diff changeset
   118
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
   119
  | vars_tr t = [t];
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
   120
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
   121
in
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
   122
72987
b1be35908165 clarified modules: avoid multiple uses of the same ML file;
wenzelm
parents: 72806
diff changeset
   123
fun hoare_vars_tr ctxt [vars, pre, prg, post] =
42153
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
   124
      let val xs = vars_tr vars
74503
403ce50e6a2a separated commands from annotations to be able to abstract about the latter only
nipkow
parents: 72998
diff changeset
   125
          val (c',a) = com_tr ctxt prg xs
403ce50e6a2a separated commands from annotations to be able to abstract about the latter only
nipkow
parents: 72998
diff changeset
   126
      in syntax_const ctxt #Valid $ assn_tr pre xs $ c' $ a $ assn_tr post xs end
72987
b1be35908165 clarified modules: avoid multiple uses of the same ML file;
wenzelm
parents: 72806
diff changeset
   127
  | hoare_vars_tr _ ts = raise TERM ("hoare_vars_tr", ts);
42153
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
   128
72987
b1be35908165 clarified modules: avoid multiple uses of the same ML file;
wenzelm
parents: 72806
diff changeset
   129
fun hoare_tc_vars_tr ctxt [vars, pre, prg, post] =
72806
4fa08e083865 Extension of session HOL/Hoare with total correctness proof system by Walter Guttmann
nipkow
parents: 69597
diff changeset
   130
      let val xs = vars_tr vars
74503
403ce50e6a2a separated commands from annotations to be able to abstract about the latter only
nipkow
parents: 72998
diff changeset
   131
          val (c',a) = com_tr ctxt prg xs
403ce50e6a2a separated commands from annotations to be able to abstract about the latter only
nipkow
parents: 72998
diff changeset
   132
      in syntax_const ctxt #ValidTC $ assn_tr pre xs $ c' $ a $ assn_tr post xs end
72987
b1be35908165 clarified modules: avoid multiple uses of the same ML file;
wenzelm
parents: 72806
diff changeset
   133
  | hoare_tc_vars_tr _ ts = raise TERM ("hoare_tc_vars_tr", ts);
72806
4fa08e083865 Extension of session HOL/Hoare with total correctness proof system by Walter Guttmann
nipkow
parents: 69597
diff changeset
   134
42153
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
   135
end;
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
   136
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
   137
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
   138
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
   139
(** print translation **)
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
   140
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
   141
local
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
   142
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
   143
fun dest_abstuple
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 61424
diff changeset
   144
      (Const (\<^const_syntax>\<open>case_prod\<close>, _) $ Abs (v, _, body)) =
42153
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
   145
        subst_bound (Syntax.free v, dest_abstuple body)
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
   146
  | dest_abstuple (Abs (v,_, body)) = subst_bound (Syntax.free v, body)
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
   147
  | dest_abstuple tm = tm;
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
   148
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 61424
diff changeset
   149
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
   150
  | abs2list (Abs (x, T, _)) = [Free (x, T)]
42153
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
   151
  | abs2list _ = [];
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
   152
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 61424
diff changeset
   153
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
   154
  | mk_ts (Abs (_, _, t)) = mk_ts t
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 61424
diff changeset
   155
  | 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
   156
  | mk_ts t = [t];
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
   157
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 61424
diff changeset
   158
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
   159
      (Syntax.free x :: abs2list t, mk_ts t)
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
   160
  | mk_vts (Abs (x, _, t)) = ([Syntax.free x], [t])
55659
4089f6e98ab9 reduced ML warnings;
wenzelm
parents: 55414
diff changeset
   161
  | mk_vts _ = raise Match;
42153
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
   162
55659
4089f6e98ab9 reduced ML warnings;
wenzelm
parents: 55414
diff changeset
   163
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
   164
  | find_ch ((v, t) :: vts) i xs =
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
   165
      if t = Bound i then find_ch vts (i - 1) xs
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
   166
      else (true, (v, subst_bounds (xs, t)));
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
   167
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 61424
diff changeset
   168
fun is_f (Const (\<^const_syntax>\<open>case_prod\<close>, _) $ Abs _) = true
55659
4089f6e98ab9 reduced ML warnings;
wenzelm
parents: 55414
diff changeset
   169
  | is_f (Abs _) = true
4089f6e98ab9 reduced ML warnings;
wenzelm
parents: 55414
diff changeset
   170
  | is_f _ = false;
42153
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
   171
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
   172
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
   173
(* assn_tr' & bexp_tr'*)
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
   174
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 61424
diff changeset
   175
fun assn_tr' (Const (\<^const_syntax>\<open>Collect\<close>, _) $ T) = dest_abstuple T
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 61424
diff changeset
   176
  | assn_tr' (Const (\<^const_syntax>\<open>inter\<close>, _) $
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 61424
diff changeset
   177
        (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
   178
      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
   179
  | assn_tr' t = t;
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
   180
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 61424
diff changeset
   181
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
   182
  | bexp_tr' t = t;
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
   183
74503
403ce50e6a2a separated commands from annotations to be able to abstract about the latter only
nipkow
parents: 72998
diff changeset
   184
fun var_tr' xo T =
403ce50e6a2a separated commands from annotations to be able to abstract about the latter only
nipkow
parents: 72998
diff changeset
   185
  let val n = dest_abstuple T
403ce50e6a2a separated commands from annotations to be able to abstract about the latter only
nipkow
parents: 72998
diff changeset
   186
  in case xo of NONE => n | SOME x => Syntax.const \<^const_syntax>\<open>HOL.eq\<close> $ Syntax.free x $ n end
72806
4fa08e083865 Extension of session HOL/Hoare with total correctness proof system by Walter Guttmann
nipkow
parents: 69597
diff changeset
   187
42153
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
   188
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
   189
(* com_tr' *)
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
   190
72987
b1be35908165 clarified modules: avoid multiple uses of the same ML file;
wenzelm
parents: 72806
diff changeset
   191
fun mk_assign ctxt f =
42153
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
   192
  let
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
   193
    val (vs, ts) = mk_vts f;
72987
b1be35908165 clarified modules: avoid multiple uses of the same ML file;
wenzelm
parents: 72806
diff changeset
   194
    val (ch, (a, b)) = find_ch (vs ~~ ts) (length vs - 1) (rev vs);
42153
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
   195
  in
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
   196
    if ch
72987
b1be35908165 clarified modules: avoid multiple uses of the same ML file;
wenzelm
parents: 72806
diff changeset
   197
    then Syntax.const \<^syntax_const>\<open>_assign\<close> $ a $ b
b1be35908165 clarified modules: avoid multiple uses of the same ML file;
wenzelm
parents: 72806
diff changeset
   198
    else syntax_const ctxt #Skip
42153
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
   199
  end;
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
   200
74503
403ce50e6a2a separated commands from annotations to be able to abstract about the latter only
nipkow
parents: 72998
diff changeset
   201
fun com_tr' ctxt (t,a) =
72987
b1be35908165 clarified modules: avoid multiple uses of the same ML file;
wenzelm
parents: 72806
diff changeset
   202
  let
b1be35908165 clarified modules: avoid multiple uses of the same ML file;
wenzelm
parents: 72806
diff changeset
   203
    val (head, args) = apfst (try (#1 o Term.dest_Const)) (Term.strip_comb t);
b1be35908165 clarified modules: avoid multiple uses of the same ML file;
wenzelm
parents: 72806
diff changeset
   204
    fun arg k = nth args (k - 1);
b1be35908165 clarified modules: avoid multiple uses of the same ML file;
wenzelm
parents: 72806
diff changeset
   205
    val n = length args;
74503
403ce50e6a2a separated commands from annotations to be able to abstract about the latter only
nipkow
parents: 72998
diff changeset
   206
    val (_, args_a) = apfst (try (#1 o Term.dest_Const)) (Term.strip_comb a);
403ce50e6a2a separated commands from annotations to be able to abstract about the latter only
nipkow
parents: 72998
diff changeset
   207
    fun arg_a k = nth args_a (k - 1)
72987
b1be35908165 clarified modules: avoid multiple uses of the same ML file;
wenzelm
parents: 72806
diff changeset
   208
  in
74503
403ce50e6a2a separated commands from annotations to be able to abstract about the latter only
nipkow
parents: 72998
diff changeset
   209
    case head of
72987
b1be35908165 clarified modules: avoid multiple uses of the same ML file;
wenzelm
parents: 72806
diff changeset
   210
      NONE => t
b1be35908165 clarified modules: avoid multiple uses of the same ML file;
wenzelm
parents: 72806
diff changeset
   211
    | SOME c =>
b1be35908165 clarified modules: avoid multiple uses of the same ML file;
wenzelm
parents: 72806
diff changeset
   212
        if c = const_syntax ctxt #Basic andalso n = 1 andalso is_f (arg 1) then
b1be35908165 clarified modules: avoid multiple uses of the same ML file;
wenzelm
parents: 72806
diff changeset
   213
          mk_assign ctxt (arg 1)
b1be35908165 clarified modules: avoid multiple uses of the same ML file;
wenzelm
parents: 72806
diff changeset
   214
        else if c = const_syntax ctxt #Seq andalso n = 2 then
74503
403ce50e6a2a separated commands from annotations to be able to abstract about the latter only
nipkow
parents: 72998
diff changeset
   215
          Syntax.const \<^syntax_const>\<open>_Seq\<close>
403ce50e6a2a separated commands from annotations to be able to abstract about the latter only
nipkow
parents: 72998
diff changeset
   216
            $ com_tr' ctxt (arg 1, arg_a 1) $ com_tr' ctxt (arg 2, arg_a 2)
72987
b1be35908165 clarified modules: avoid multiple uses of the same ML file;
wenzelm
parents: 72806
diff changeset
   217
        else if c = const_syntax ctxt #Cond andalso n = 3 then
b1be35908165 clarified modules: avoid multiple uses of the same ML file;
wenzelm
parents: 72806
diff changeset
   218
          Syntax.const \<^syntax_const>\<open>_Cond\<close> $
74503
403ce50e6a2a separated commands from annotations to be able to abstract about the latter only
nipkow
parents: 72998
diff changeset
   219
            bexp_tr' (arg 1) $ com_tr' ctxt (arg 2, arg_a 1) $ com_tr' ctxt (arg 3, arg_a 2)
403ce50e6a2a separated commands from annotations to be able to abstract about the latter only
nipkow
parents: 72998
diff changeset
   220
        else if c = const_syntax ctxt #While andalso n = 2 then
403ce50e6a2a separated commands from annotations to be able to abstract about the latter only
nipkow
parents: 72998
diff changeset
   221
               let val (xo,a') = case arg_a 3 of
403ce50e6a2a separated commands from annotations to be able to abstract about the latter only
nipkow
parents: 72998
diff changeset
   222
                     Abs(x,_,a) => (if loose_bvar1(a,0) then SOME x else NONE,
403ce50e6a2a separated commands from annotations to be able to abstract about the latter only
nipkow
parents: 72998
diff changeset
   223
                                    subst_bound (Syntax.free x, a)) |
403ce50e6a2a separated commands from annotations to be able to abstract about the latter only
nipkow
parents: 72998
diff changeset
   224
                     t => (NONE,t)
403ce50e6a2a separated commands from annotations to be able to abstract about the latter only
nipkow
parents: 72998
diff changeset
   225
               in Syntax.const \<^syntax_const>\<open>_While\<close>
403ce50e6a2a separated commands from annotations to be able to abstract about the latter only
nipkow
parents: 72998
diff changeset
   226
                 $ bexp_tr' (arg 1) $ assn_tr' (arg_a 1) $ var_tr' xo (arg_a 2) $ com_tr' ctxt (arg 2, a')
403ce50e6a2a separated commands from annotations to be able to abstract about the latter only
nipkow
parents: 72998
diff changeset
   227
               end
403ce50e6a2a separated commands from annotations to be able to abstract about the latter only
nipkow
parents: 72998
diff changeset
   228
        else t
72987
b1be35908165 clarified modules: avoid multiple uses of the same ML file;
wenzelm
parents: 72806
diff changeset
   229
  end;
42153
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
   230
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
   231
in
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
   232
74503
403ce50e6a2a separated commands from annotations to be able to abstract about the latter only
nipkow
parents: 72998
diff changeset
   233
fun spec_tr' syn ctxt [p, c, a, q] =
403ce50e6a2a separated commands from annotations to be able to abstract about the latter only
nipkow
parents: 72998
diff changeset
   234
  Syntax.const syn $ assn_tr' p $ com_tr' ctxt (c,a) $ assn_tr' q;
42153
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
   235
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
   236
end;
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
   237
72987
b1be35908165 clarified modules: avoid multiple uses of the same ML file;
wenzelm
parents: 72806
diff changeset
   238
b1be35908165 clarified modules: avoid multiple uses of the same ML file;
wenzelm
parents: 72806
diff changeset
   239
(** setup **)
b1be35908165 clarified modules: avoid multiple uses of the same ML file;
wenzelm
parents: 72806
diff changeset
   240
72998
7ea253f93606 more robust: clarified error when merging Hoare_Logic + Hoare_Logic_Abort;
wenzelm
parents: 72987
diff changeset
   241
val _ =
7ea253f93606 more robust: clarified error when merging Hoare_Logic + Hoare_Logic_Abort;
wenzelm
parents: 72987
diff changeset
   242
  Theory.setup
7ea253f93606 more robust: clarified error when merging Hoare_Logic + Hoare_Logic_Abort;
wenzelm
parents: 72987
diff changeset
   243
   (Sign.parse_translation
7ea253f93606 more robust: clarified error when merging Hoare_Logic + Hoare_Logic_Abort;
wenzelm
parents: 72987
diff changeset
   244
    [(\<^syntax_const>\<open>_hoare_vars\<close>, hoare_vars_tr),
7ea253f93606 more robust: clarified error when merging Hoare_Logic + Hoare_Logic_Abort;
wenzelm
parents: 72987
diff changeset
   245
     (\<^syntax_const>\<open>_hoare_vars_tc\<close>, hoare_tc_vars_tr)]);
7ea253f93606 more robust: clarified error when merging Hoare_Logic + Hoare_Logic_Abort;
wenzelm
parents: 72987
diff changeset
   246
72987
b1be35908165 clarified modules: avoid multiple uses of the same ML file;
wenzelm
parents: 72806
diff changeset
   247
fun setup consts =
b1be35908165 clarified modules: avoid multiple uses of the same ML file;
wenzelm
parents: 72806
diff changeset
   248
  Data.put (SOME consts) #>
b1be35908165 clarified modules: avoid multiple uses of the same ML file;
wenzelm
parents: 72806
diff changeset
   249
  Sign.print_translation
b1be35908165 clarified modules: avoid multiple uses of the same ML file;
wenzelm
parents: 72806
diff changeset
   250
   [(#Valid consts, spec_tr' \<^syntax_const>\<open>_hoare\<close>),
b1be35908165 clarified modules: avoid multiple uses of the same ML file;
wenzelm
parents: 72806
diff changeset
   251
    (#ValidTC consts, spec_tr' \<^syntax_const>\<open>_hoare_tc\<close>)];
b1be35908165 clarified modules: avoid multiple uses of the same ML file;
wenzelm
parents: 72806
diff changeset
   252
42153
fa108629d132 use shared copy of hoare_syntax.ML;
wenzelm
parents:
diff changeset
   253
end;