src/HOL/IMPP/Hoare.thy
author wenzelm
Fri, 05 May 2006 21:59:41 +0200
changeset 19573 340c466c9605
parent 17477 ceb42ea2f223
child 19803 aa2581752afb
permissions -rw-r--r--
axiomatization;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
8177
e59e93ad85eb added IMPP to HOL
oheimb
parents:
diff changeset
     1
(*  Title:      HOL/IMPP/Hoare.thy
e59e93ad85eb added IMPP to HOL
oheimb
parents:
diff changeset
     2
    ID:         $Id$
e59e93ad85eb added IMPP to HOL
oheimb
parents:
diff changeset
     3
    Author:     David von Oheimb
e59e93ad85eb added IMPP to HOL
oheimb
parents:
diff changeset
     4
    Copyright   1999 TUM
e59e93ad85eb added IMPP to HOL
oheimb
parents:
diff changeset
     5
*)
e59e93ad85eb added IMPP to HOL
oheimb
parents:
diff changeset
     6
17477
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 10834
diff changeset
     7
header {* Inductive definition of Hoare logic for partial correctness *}
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 10834
diff changeset
     8
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 10834
diff changeset
     9
theory Hoare
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 10834
diff changeset
    10
imports Natural
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 10834
diff changeset
    11
begin
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 10834
diff changeset
    12
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 10834
diff changeset
    13
text {*
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 10834
diff changeset
    14
  Completeness is taken relative to completeness of the underlying logic.
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 10834
diff changeset
    15
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 10834
diff changeset
    16
  Two versions of completeness proof: nested single recursion
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 10834
diff changeset
    17
  vs. simultaneous recursion in call rule
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 10834
diff changeset
    18
*}
8177
e59e93ad85eb added IMPP to HOL
oheimb
parents:
diff changeset
    19
e59e93ad85eb added IMPP to HOL
oheimb
parents:
diff changeset
    20
types 'a assn = "'a => state => bool"
e59e93ad85eb added IMPP to HOL
oheimb
parents:
diff changeset
    21
translations
17477
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 10834
diff changeset
    22
  "a assn"   <= (type)"a => state => bool"
8177
e59e93ad85eb added IMPP to HOL
oheimb
parents:
diff changeset
    23
e59e93ad85eb added IMPP to HOL
oheimb
parents:
diff changeset
    24
constdefs
e59e93ad85eb added IMPP to HOL
oheimb
parents:
diff changeset
    25
  state_not_singleton :: bool
17477
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 10834
diff changeset
    26
  "state_not_singleton == \<exists>s t::state. s ~= t" (* at least two elements *)
8177
e59e93ad85eb added IMPP to HOL
oheimb
parents:
diff changeset
    27
e59e93ad85eb added IMPP to HOL
oheimb
parents:
diff changeset
    28
  peek_and    :: "'a assn => (state => bool) => 'a assn" (infixr "&>" 35)
17477
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 10834
diff changeset
    29
  "peek_and P p == %Z s. P Z s & p s"
8177
e59e93ad85eb added IMPP to HOL
oheimb
parents:
diff changeset
    30
e59e93ad85eb added IMPP to HOL
oheimb
parents:
diff changeset
    31
datatype 'a triple =
17477
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 10834
diff changeset
    32
  triple "'a assn"  com  "'a assn"       ("{(1_)}./ (_)/ .{(1_)}" [3,60,3] 58)
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 10834
diff changeset
    33
8177
e59e93ad85eb added IMPP to HOL
oheimb
parents:
diff changeset
    34
consts
17477
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 10834
diff changeset
    35
  triple_valid ::            "nat => 'a triple     => bool" ( "|=_:_" [0 , 58] 57)
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 10834
diff changeset
    36
  hoare_valids ::  "'a triple set => 'a triple set => bool" ("_||=_"  [58, 58] 57)
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 10834
diff changeset
    37
  hoare_derivs :: "('a triple set *  'a triple set)   set"
8177
e59e93ad85eb added IMPP to HOL
oheimb
parents:
diff changeset
    38
syntax
17477
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 10834
diff changeset
    39
  triples_valid::            "nat => 'a triple set => bool" ("||=_:_" [0 , 58] 57)
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 10834
diff changeset
    40
  hoare_valid  ::  "'a triple set => 'a triple     => bool" ("_|=_"   [58, 58] 57)
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 10834
diff changeset
    41
"@hoare_derivs"::  "'a triple set => 'a triple set => bool" ("_||-_"  [58, 58] 57)
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 10834
diff changeset
    42
"@hoare_deriv" ::  "'a triple set => 'a triple     => bool" ("_|-_"   [58, 58] 57)
8177
e59e93ad85eb added IMPP to HOL
oheimb
parents:
diff changeset
    43
17477
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 10834
diff changeset
    44
defs triple_valid_def: "|=n:t  ==  case t of {P}.c.{Q} =>
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 10834
diff changeset
    45
                                !Z s. P Z s --> (!s'. <c,s> -n-> s' --> Q Z s')"
8177
e59e93ad85eb added IMPP to HOL
oheimb
parents:
diff changeset
    46
translations          "||=n:G" == "Ball G (triple_valid n)"
17477
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 10834
diff changeset
    47
defs hoare_valids_def: "G||=ts   ==  !n. ||=n:G --> ||=n:ts"
8177
e59e93ad85eb added IMPP to HOL
oheimb
parents:
diff changeset
    48
translations         "G |=t  " == " G||={t}"
e59e93ad85eb added IMPP to HOL
oheimb
parents:
diff changeset
    49
                     "G||-ts"  == "(G,ts) : hoare_derivs"
e59e93ad85eb added IMPP to HOL
oheimb
parents:
diff changeset
    50
                     "G |-t"   == " G||-{t}"
e59e93ad85eb added IMPP to HOL
oheimb
parents:
diff changeset
    51
e59e93ad85eb added IMPP to HOL
oheimb
parents:
diff changeset
    52
(* Most General Triples *)
17477
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 10834
diff changeset
    53
constdefs MGT    :: "com => state triple"            ("{=}._.{->}" [60] 58)
8177
e59e93ad85eb added IMPP to HOL
oheimb
parents:
diff changeset
    54
         "{=}.c.{->} == {%Z s0. Z = s0}. c .{%Z s1. <c,Z> -c-> s1}"
e59e93ad85eb added IMPP to HOL
oheimb
parents:
diff changeset
    55
17477
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 10834
diff changeset
    56
inductive hoare_derivs intros
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 10834
diff changeset
    57
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 10834
diff changeset
    58
  empty:    "G||-{}"
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 10834
diff changeset
    59
  insert: "[| G |-t;  G||-ts |]
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 10834
diff changeset
    60
        ==> G||-insert t ts"
8177
e59e93ad85eb added IMPP to HOL
oheimb
parents:
diff changeset
    61
17477
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 10834
diff changeset
    62
  asm:      "ts <= G ==>
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 10834
diff changeset
    63
             G||-ts" (* {P}.BODY pn.{Q} instead of (general) t for SkipD_lemma *)
8177
e59e93ad85eb added IMPP to HOL
oheimb
parents:
diff changeset
    64
17477
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 10834
diff changeset
    65
  cut:   "[| G'||-ts; G||-G' |] ==> G||-ts" (* for convenience and efficiency *)
8177
e59e93ad85eb added IMPP to HOL
oheimb
parents:
diff changeset
    66
17477
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 10834
diff changeset
    67
  weaken: "[| G||-ts' ; ts <= ts' |] ==> G||-ts"
8177
e59e93ad85eb added IMPP to HOL
oheimb
parents:
diff changeset
    68
17477
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 10834
diff changeset
    69
  conseq: "!Z s. P  Z  s --> (? P' Q'. G|-{P'}.c.{Q'} &
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 10834
diff changeset
    70
                                   (!s'. (!Z'. P' Z' s --> Q' Z' s') --> Q Z s'))
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 10834
diff changeset
    71
          ==> G|-{P}.c.{Q}"
8177
e59e93ad85eb added IMPP to HOL
oheimb
parents:
diff changeset
    72
e59e93ad85eb added IMPP to HOL
oheimb
parents:
diff changeset
    73
17477
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 10834
diff changeset
    74
  Skip:  "G|-{P}. SKIP .{P}"
8177
e59e93ad85eb added IMPP to HOL
oheimb
parents:
diff changeset
    75
17477
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 10834
diff changeset
    76
  Ass:   "G|-{%Z s. P Z (s[X::=a s])}. X:==a .{P}"
8177
e59e93ad85eb added IMPP to HOL
oheimb
parents:
diff changeset
    77
17477
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 10834
diff changeset
    78
  Local: "G|-{P}. c .{%Z s. Q Z (s[Loc X::=s'<X>])}
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 10834
diff changeset
    79
      ==> G|-{%Z s. s'=s & P Z (s[Loc X::=a s])}. LOCAL X:=a IN c .{Q}"
8177
e59e93ad85eb added IMPP to HOL
oheimb
parents:
diff changeset
    80
17477
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 10834
diff changeset
    81
  Comp:  "[| G|-{P}.c.{Q};
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 10834
diff changeset
    82
             G|-{Q}.d.{R} |]
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 10834
diff changeset
    83
         ==> G|-{P}. (c;;d) .{R}"
8177
e59e93ad85eb added IMPP to HOL
oheimb
parents:
diff changeset
    84
17477
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 10834
diff changeset
    85
  If:    "[| G|-{P &>        b }.c.{Q};
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 10834
diff changeset
    86
             G|-{P &> (Not o b)}.d.{Q} |]
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 10834
diff changeset
    87
         ==> G|-{P}. IF b THEN c ELSE d .{Q}"
8177
e59e93ad85eb added IMPP to HOL
oheimb
parents:
diff changeset
    88
17477
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 10834
diff changeset
    89
  Loop:  "G|-{P &> b}.c.{P} ==>
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 10834
diff changeset
    90
          G|-{P}. WHILE b DO c .{P &> (Not o b)}"
8177
e59e93ad85eb added IMPP to HOL
oheimb
parents:
diff changeset
    91
e59e93ad85eb added IMPP to HOL
oheimb
parents:
diff changeset
    92
(*
17477
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 10834
diff changeset
    93
  BodyN: "(insert ({P}. BODY pn  .{Q}) G)
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 10834
diff changeset
    94
           |-{P}.  the (body pn) .{Q} ==>
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 10834
diff changeset
    95
          G|-{P}.       BODY pn  .{Q}"
8177
e59e93ad85eb added IMPP to HOL
oheimb
parents:
diff changeset
    96
*)
17477
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 10834
diff changeset
    97
  Body:  "[| G Un (%p. {P p}.      BODY p  .{Q p})`Procs
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 10834
diff changeset
    98
               ||-(%p. {P p}. the (body p) .{Q p})`Procs |]
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 10834
diff changeset
    99
         ==>  G||-(%p. {P p}.      BODY p  .{Q p})`Procs"
8177
e59e93ad85eb added IMPP to HOL
oheimb
parents:
diff changeset
   100
17477
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 10834
diff changeset
   101
  Call:     "G|-{P}. BODY pn .{%Z s. Q Z (setlocs s (getlocs s')[X::=s<Res>])}
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 10834
diff changeset
   102
         ==> G|-{%Z s. s'=s & P Z (setlocs s newlocs[Loc Arg::=a s])}.
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 10834
diff changeset
   103
             X:=CALL pn(a) .{Q}"
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 10834
diff changeset
   104
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 10834
diff changeset
   105
ML {* use_legacy_bindings (the_context ()) *}
8177
e59e93ad85eb added IMPP to HOL
oheimb
parents:
diff changeset
   106
e59e93ad85eb added IMPP to HOL
oheimb
parents:
diff changeset
   107
end