src/HOL/Hoare/HoareAbort.thy
changeset 35054 a5db9779b026
parent 34940 3e80eab831a1
child 35101 6ce9177d6b38
equal deleted inserted replaced
35053:43175817d83b 35054:a5db9779b026
    19    | Abort
    19    | Abort
    20    | Seq "'a com" "'a com"               ("(_;/ _)"      [61,60] 60)
    20    | Seq "'a com" "'a com"               ("(_;/ _)"      [61,60] 60)
    21    | Cond "'a bexp" "'a com" "'a com"    ("(1IF _/ THEN _ / ELSE _/ FI)"  [0,0,0] 61)
    21    | Cond "'a bexp" "'a com" "'a com"    ("(1IF _/ THEN _ / ELSE _/ FI)"  [0,0,0] 61)
    22    | While "'a bexp" "'a assn" "'a com"  ("(1WHILE _/ INV {_} //DO _ /OD)"  [0,0,0] 61)
    22    | While "'a bexp" "'a assn" "'a com"  ("(1WHILE _/ INV {_} //DO _ /OD)"  [0,0,0] 61)
    23   
    23   
    24 syntax
    24 abbreviation annskip ("SKIP") where "SKIP == Basic id"
    25   "@assign"  :: "id => 'b => 'a com"        ("(2_ :=/ _)" [70,65] 61)
       
    26   "@annskip" :: "'a com"                    ("SKIP")
       
    27 
       
    28 translations
       
    29             "SKIP" == "Basic id"
       
    30 
    25 
    31 types 'a sem = "'a option => 'a option => bool"
    26 types 'a sem = "'a option => 'a option => bool"
    32 
    27 
    33 consts iter :: "nat => 'a bexp => 'a sem => 'a sem"
    28 consts iter :: "nat => 'a bexp => 'a sem => 'a sem"
    34 primrec
    29 primrec
    49 
    44 
    50 constdefs Valid :: "'a bexp \<Rightarrow> 'a com \<Rightarrow> 'a bexp \<Rightarrow> bool"
    45 constdefs Valid :: "'a bexp \<Rightarrow> 'a com \<Rightarrow> 'a bexp \<Rightarrow> bool"
    51   "Valid p c q == \<forall>s s'. Sem c s s' \<longrightarrow> s : Some ` p \<longrightarrow> s' : Some ` q"
    46   "Valid p c q == \<forall>s s'. Sem c s s' \<longrightarrow> s : Some ` p \<longrightarrow> s' : Some ` q"
    52 
    47 
    53 
    48 
       
    49 
       
    50 (** parse translations **)
       
    51 
    54 syntax
    52 syntax
    55  "@hoare_vars" :: "[idts, 'a assn,'a com,'a assn] => bool"
    53   "_assign"  :: "id => 'b => 'a com"        ("(2_ :=/ _)" [70,65] 61)
       
    54 
       
    55 syntax
       
    56   "_hoare_vars" :: "[idts, 'a assn,'a com,'a assn] => bool"
    56                  ("VARS _// {_} // _ // {_}" [0,0,55,0] 50)
    57                  ("VARS _// {_} // _ // {_}" [0,0,55,0] 50)
    57 syntax ("" output)
    58 syntax ("" output)
    58  "@hoare"      :: "['a assn,'a com,'a assn] => bool"
    59   "_hoare"      :: "['a assn,'a com,'a assn] => bool"
    59                  ("{_} // _ // {_}" [0,55,0] 50)
    60                  ("{_} // _ // {_}" [0,55,0] 50)
    60 
    61 ML {*
    61 (** parse translations **)
       
    62 
       
    63 ML{*
       
    64 
    62 
    65 local
    63 local
    66 fun free a = Free(a,dummyT)
    64 fun free a = Free(a,dummyT)
    67 fun abs((a,T),body) =
    65 fun abs((a,T),body) =
    68   let val a = absfree(a, dummyT, body)
    66   let val a = absfree(a, dummyT, body)
    90   
    88   
    91 fun assn_tr r xs = Syntax.const "Collect" $ mk_abstuple xs r;
    89 fun assn_tr r xs = Syntax.const "Collect" $ mk_abstuple xs r;
    92 *}
    90 *}
    93 (* com_tr *)
    91 (* com_tr *)
    94 ML{*
    92 ML{*
    95 fun com_tr (Const("@assign",_) $ Free (a,_) $ e) xs =
    93 fun com_tr (Const("_assign",_) $ Free (a,_) $ e) xs =
    96       Syntax.const "Basic" $ mk_fexp a e xs
    94       Syntax.const "Basic" $ mk_fexp a e xs
    97   | com_tr (Const ("Basic",_) $ f) xs = Syntax.const "Basic" $ f
    95   | com_tr (Const ("Basic",_) $ f) xs = Syntax.const "Basic" $ f
    98   | com_tr (Const ("Seq",_) $ c1 $ c2) xs =
    96   | com_tr (Const ("Seq",_) $ c1 $ c2) xs =
    99       Syntax.const "Seq" $ com_tr c1 xs $ com_tr c2 xs
    97       Syntax.const "Seq" $ com_tr c1 xs $ com_tr c2 xs
   100   | com_tr (Const ("Cond",_) $ b $ c1 $ c2) xs =
    98   | com_tr (Const ("Cond",_) $ b $ c1 $ c2) xs =
   122       end
   120       end
   123   | hoare_vars_tr ts = raise TERM ("hoare_vars_tr", ts);
   121   | hoare_vars_tr ts = raise TERM ("hoare_vars_tr", ts);
   124 end
   122 end
   125 *}
   123 *}
   126 
   124 
   127 parse_translation {* [("@hoare_vars", hoare_vars_tr)] *}
   125 parse_translation {* [("_hoare_vars", hoare_vars_tr)] *}
   128 
   126 
   129 
   127 
   130 (*****************************************************************************)
   128 (*****************************************************************************)
   131 
   129 
   132 (*** print translations ***)
   130 (*** print translations ***)
   174 (*com_tr' *)
   172 (*com_tr' *)
   175 ML{*
   173 ML{*
   176 fun mk_assign f =
   174 fun mk_assign f =
   177   let val (vs, ts) = mk_vts f;
   175   let val (vs, ts) = mk_vts f;
   178       val (ch, which) = find_ch (vs~~ts) ((length vs)-1) (rev vs)
   176       val (ch, which) = find_ch (vs~~ts) ((length vs)-1) (rev vs)
   179   in if ch then Syntax.const "@assign" $ fst(which) $ snd(which)
   177   in if ch then Syntax.const "_assign" $ fst(which) $ snd(which)
   180      else Syntax.const "@skip" end;
   178      else Syntax.const @{const_syntax annskip} end;
   181 
   179 
   182 fun com_tr' (Const ("Basic",_) $ f) = if is_f f then mk_assign f
   180 fun com_tr' (Const ("Basic",_) $ f) = if is_f f then mk_assign f
   183                                            else Syntax.const "Basic" $ f
   181                                            else Syntax.const "Basic" $ f
   184   | com_tr' (Const ("Seq",_) $ c1 $ c2) = Syntax.const "Seq" $
   182   | com_tr' (Const ("Seq",_) $ c1 $ c2) = Syntax.const "Seq" $
   185                                                  com_tr' c1 $ com_tr' c2
   183                                                  com_tr' c1 $ com_tr' c2
   189                                                bexp_tr' b $ assn_tr' I $ com_tr' c
   187                                                bexp_tr' b $ assn_tr' I $ com_tr' c
   190   | com_tr' t = t;
   188   | com_tr' t = t;
   191 
   189 
   192 
   190 
   193 fun spec_tr' [p, c, q] =
   191 fun spec_tr' [p, c, q] =
   194   Syntax.const "@hoare" $ assn_tr' p $ com_tr' c $ assn_tr' q
   192   Syntax.const "_hoare" $ assn_tr' p $ com_tr' c $ assn_tr' q
   195 *}
   193 *}
   196 
   194 
   197 print_translation {* [("Valid", spec_tr')] *}
   195 print_translation {* [(@{const_syntax Valid}, spec_tr')] *}
   198 
   196 
   199 (*** The proof rules ***)
   197 (*** The proof rules ***)
   200 
   198 
   201 lemma SkipRule: "p \<subseteq> q \<Longrightarrow> Valid p (Basic id) q"
   199 lemma SkipRule: "p \<subseteq> q \<Longrightarrow> Valid p (Basic id) q"
   202 by (auto simp:Valid_def)
   200 by (auto simp:Valid_def)
   255 
   253 
   256 (* Special syntax for guarded statements and guarded array updates: *)
   254 (* Special syntax for guarded statements and guarded array updates: *)
   257 
   255 
   258 syntax
   256 syntax
   259   guarded_com :: "bool \<Rightarrow> 'a com \<Rightarrow> 'a com"  ("(2_ \<rightarrow>/ _)" 71)
   257   guarded_com :: "bool \<Rightarrow> 'a com \<Rightarrow> 'a com"  ("(2_ \<rightarrow>/ _)" 71)
   260   array_update :: "'a list \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> 'a com"  ("(2_[_] :=/ _)" [70,65] 61)
   258   array_update :: "'a list \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> 'a com"  ("(2_[_] :=/ _)" [70, 65] 61)
   261 translations
   259 translations
   262   "P \<rightarrow> c" == "IF P THEN c ELSE Abort FI"
   260   "P \<rightarrow> c" == "IF P THEN c ELSE Abort FI"
   263   "a[i] := v" => "(i < CONST length a) \<rightarrow> (a := CONST list_update a i v)"
   261   "a[i] := v" => "(i < CONST length a) \<rightarrow> (a := CONST list_update a i v)"
   264   (* reverse translation not possible because of duplicate "a" *)
   262   (* reverse translation not possible because of duplicate "a" *)
   265 
   263