src/HOL/Hoare/Hoare_Logic_Abort.thy
author wenzelm
Sat, 25 May 2013 15:37:53 +0200
changeset 53280 36ffe23b25f8
parent 52854 9e7d1c139569
child 57002 f0f895716a8b
permissions -rw-r--r--
syntax translations always depend on context;
     1 (*  Title:      HOL/Hoare/Hoare_Logic_Abort.thy
     2     Author:     Leonor Prensa Nieto & Tobias Nipkow
     3     Copyright   2003 TUM
     4 
     5 Like Hoare.thy, but with an Abort statement for modelling run time errors.
     6 *)
     7 
     8 theory Hoare_Logic_Abort
     9 imports Main
    10 begin
    11 
    12 type_synonym 'a bexp = "'a set"
    13 type_synonym 'a assn = "'a set"
    14 
    15 datatype
    16  'a com = Basic "'a \<Rightarrow> 'a"
    17    | Abort
    18    | Seq "'a com" "'a com"               ("(_;/ _)"      [61,60] 60)
    19    | Cond "'a bexp" "'a com" "'a com"    ("(1IF _/ THEN _ / ELSE _/ FI)"  [0,0,0] 61)
    20    | While "'a bexp" "'a assn" "'a com"  ("(1WHILE _/ INV {_} //DO _ /OD)"  [0,0,0] 61)
    21 
    22 abbreviation annskip ("SKIP") where "SKIP == Basic id"
    23 
    24 type_synonym 'a sem = "'a option => 'a option => bool"
    25 
    26 inductive Sem :: "'a com \<Rightarrow> 'a sem"
    27 where
    28   "Sem (Basic f) None None"
    29 | "Sem (Basic f) (Some s) (Some (f s))"
    30 | "Sem Abort s None"
    31 | "Sem c1 s s'' \<Longrightarrow> Sem c2 s'' s' \<Longrightarrow> Sem (c1;c2) s s'"
    32 | "Sem (IF b THEN c1 ELSE c2 FI) None None"
    33 | "s \<in> b \<Longrightarrow> Sem c1 (Some s) s' \<Longrightarrow> Sem (IF b THEN c1 ELSE c2 FI) (Some s) s'"
    34 | "s \<notin> b \<Longrightarrow> Sem c2 (Some s) s' \<Longrightarrow> Sem (IF b THEN c1 ELSE c2 FI) (Some s) s'"
    35 | "Sem (While b x c) None None"
    36 | "s \<notin> b \<Longrightarrow> Sem (While b x c) (Some s) (Some s)"
    37 | "s \<in> b \<Longrightarrow> Sem c (Some s) s'' \<Longrightarrow> Sem (While b x c) s'' s' \<Longrightarrow>
    38    Sem (While b x c) (Some s) s'"
    39 
    40 inductive_cases [elim!]:
    41   "Sem (Basic f) s s'" "Sem (c1;c2) s s'"
    42   "Sem (IF b THEN c1 ELSE c2 FI) s s'"
    43 
    44 definition Valid :: "'a bexp \<Rightarrow> 'a com \<Rightarrow> 'a bexp \<Rightarrow> bool" where
    45   "Valid p c q == \<forall>s s'. Sem c s s' \<longrightarrow> s : Some ` p \<longrightarrow> s' : Some ` q"
    46 
    47 
    48 syntax
    49   "_assign" :: "idt => 'b => 'a com"  ("(2_ :=/ _)" [70, 65] 61)
    50 
    51 syntax
    52   "_hoare_abort_vars" :: "[idts, 'a assn,'a com,'a assn] => bool"
    53                  ("VARS _// {_} // _ // {_}" [0,0,55,0] 50)
    54 syntax ("" output)
    55   "_hoare_abort"      :: "['a assn,'a com,'a assn] => bool"
    56                  ("{_} // _ // {_}" [0,55,0] 50)
    57 
    58 ML_file "hoare_syntax.ML"
    59 parse_translation {* [(@{syntax_const "_hoare_abort_vars"}, K Hoare_Syntax.hoare_vars_tr)] *}
    60 print_translation
    61   {* [(@{const_syntax Valid}, K (Hoare_Syntax.spec_tr' @{syntax_const "_hoare_abort"}))] *}
    62 
    63 
    64 (*** The proof rules ***)
    65 
    66 lemma SkipRule: "p \<subseteq> q \<Longrightarrow> Valid p (Basic id) q"
    67 by (auto simp:Valid_def)
    68 
    69 lemma BasicRule: "p \<subseteq> {s. f s \<in> q} \<Longrightarrow> Valid p (Basic f) q"
    70 by (auto simp:Valid_def)
    71 
    72 lemma SeqRule: "Valid P c1 Q \<Longrightarrow> Valid Q c2 R \<Longrightarrow> Valid P (c1;c2) R"
    73 by (auto simp:Valid_def)
    74 
    75 lemma CondRule:
    76  "p \<subseteq> {s. (s \<in> b \<longrightarrow> s \<in> w) \<and> (s \<notin> b \<longrightarrow> s \<in> w')}
    77   \<Longrightarrow> Valid w c1 q \<Longrightarrow> Valid w' c2 q \<Longrightarrow> Valid p (Cond b c1 c2) q"
    78 by (fastforce simp:Valid_def image_def)
    79 
    80 lemma While_aux:
    81   assumes "Sem (WHILE b INV {i} DO c OD) s s'"
    82   shows "\<forall>s s'. Sem c s s' \<longrightarrow> s \<in> Some ` (I \<inter> b) \<longrightarrow> s' \<in> Some ` I \<Longrightarrow>
    83     s \<in> Some ` I \<Longrightarrow> s' \<in> Some ` (I \<inter> -b)"
    84   using assms
    85   by (induct "WHILE b INV {i} DO c OD" s s') auto
    86 
    87 lemma WhileRule:
    88  "p \<subseteq> i \<Longrightarrow> Valid (i \<inter> b) c i \<Longrightarrow> i \<inter> (-b) \<subseteq> q \<Longrightarrow> Valid p (While b i c) q"
    89 apply(simp add:Valid_def)
    90 apply(simp (no_asm) add:image_def)
    91 apply clarify
    92 apply(drule While_aux)
    93   apply assumption
    94  apply blast
    95 apply blast
    96 done
    97 
    98 lemma AbortRule: "p \<subseteq> {s. False} \<Longrightarrow> Valid p Abort q"
    99 by(auto simp:Valid_def)
   100 
   101 
   102 subsection {* Derivation of the proof rules and, most importantly, the VCG tactic *}
   103 
   104 lemma Compl_Collect: "-(Collect b) = {x. ~(b x)}"
   105   by blast
   106 
   107 ML_file "hoare_tac.ML"
   108 
   109 method_setup vcg = {*
   110   Scan.succeed (fn ctxt => SIMPLE_METHOD' (hoare_tac ctxt (K all_tac))) *}
   111   "verification condition generator"
   112 
   113 method_setup vcg_simp = {*
   114   Scan.succeed (fn ctxt =>
   115     SIMPLE_METHOD' (hoare_tac ctxt (asm_full_simp_tac ctxt))) *}
   116   "verification condition generator plus simplification"
   117 
   118 (* Special syntax for guarded statements and guarded array updates: *)
   119 
   120 syntax
   121   "_guarded_com" :: "bool \<Rightarrow> 'a com \<Rightarrow> 'a com"  ("(2_ \<rightarrow>/ _)" 71)
   122   "_array_update" :: "'a list \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> 'a com"  ("(2_[_] :=/ _)" [70, 65] 61)
   123 translations
   124   "P \<rightarrow> c" == "IF P THEN c ELSE CONST Abort FI"
   125   "a[i] := v" => "(i < CONST length a) \<rightarrow> (a := CONST list_update a i v)"
   126   (* reverse translation not possible because of duplicate "a" *)
   127 
   128 text{* Note: there is no special syntax for guarded array access. Thus
   129 you must write @{text"j < length a \<rightarrow> a[i] := a!j"}. *}
   130 
   131 end