src/HOL/NanoJava/AxSem.thy
changeset 11497 0e66e0114d9a
parent 11486 8f32860eac3a
child 11507 4b32a46ffd29
equal deleted inserted replaced
11496:fa8d12b789e1 11497:0e66e0114d9a
    75                                  | Addr a => obj_class s a <=C C) --> Q v s} ==>
    75                                  | Addr a => obj_class s a <=C C) --> Q v s} ==>
    76          A |-e {P} Cast C e {Q}"
    76          A |-e {P} Cast C e {Q}"
    77 
    77 
    78   Call: "[| A |-e {P} e1 {Q}; \<forall>a. A |-e {Q a} e2 {R a};
    78   Call: "[| A |-e {P} e1 {Q}; \<forall>a. A |-e {Q a} e2 {R a};
    79     \<forall>a p l. A |- {\<lambda>s'. \<exists>s. R a p s \<and> l = s \<and> 
    79     \<forall>a p l. A |- {\<lambda>s'. \<exists>s. R a p s \<and> l = s \<and> 
    80                     s' = lupd(This\<mapsto>a)(lupd(Param\<mapsto>p)(init_locs C m s))}
    80                     s' = lupd(This\<mapsto>a)(lupd(Param\<mapsto>p)(reset_locs s))}
    81                   Meth C m {\<lambda>s. S (s<Res>) (set_locs l s)} |] ==>
    81                   Meth C m {\<lambda>s. S (s<Res>) (set_locs l s)} |] ==>
    82              A |-e {P} {C}e1..m(e2) {S}"
    82              A |-e {P} {C}e1..m(e2) {S}"
    83 
    83 
    84   Meth: "\<forall>D. A |- {\<lambda>s. \<exists>a. s<This> = Addr a \<and> D=obj_class s a \<and> D <=C C \<and> P s}
    84   Meth: "\<forall>D. A |- {\<lambda>s'. \<exists>s a. s<This> = Addr a \<and> D = obj_class s a \<and> D <=C C \<and> 
    85                   Impl D m {Q} ==>
    85                         P s \<and> s' = init_locs D m s}
       
    86                   Impl (D,m) {Q} ==>
    86              A |- {P} Meth C m {Q}"
    87              A |- {P} Meth C m {Q}"
    87 
    88 
    88   (*\<Union>z instead of \<forall>z in the conclusion and
    89   (*\<Union>z instead of \<forall>z in the conclusion and
    89       z restricted to type state due to limitations of the inductive package *)
    90       z restricted to type state due to limitations of the inductive package *)
    90   Impl: "A\<union>   (\<Union>z::state. (\<lambda>(C,m). (P z C m, Impl C m, Q z C m))`ms) ||- 
    91   Impl: "\<forall>z::state. A\<union> (\<Union>z. (\<lambda>M. (P z M, Impl M, Q z M))`Ms) ||- 
    91                (\<Union>z::state. (\<lambda>(C,m). (P z C m, body C m, Q z C m))`ms) ==>
    92                             (\<lambda>M. (P z M, body M, Q z M))`Ms ==>
    92          A ||- (\<Union>z::state. (\<lambda>(C,m). (P z C m, Impl C m, Q z C m))`ms)"
    93                       A ||- (\<lambda>M. (P z M, Impl M, Q z M))`Ms"
    93 
    94 
    94 (* structural rules *)
    95 (* structural rules *)
    95 
    96 
    96   Asm:  "   a \<in> A ==> A ||- {a}"
    97   Asm:  "   a \<in> A ==> A ||- {a}"
    97 
    98 
    98   ConjI: " \<forall>c \<in> C. A ||- {c} ==> A ||- C"
    99   ConjI: " \<forall>c \<in> C. A ||- {c} ==> A ||- C"
    99 
   100 
   100   ConjE: "[|A ||- C; c \<in> C |] ==> A ||- {c}"
   101   ConjE: "[|A ||- C; c \<in> C |] ==> A ||- {c}"
   101 
   102 
   102   (* z restricted to type state due to limitations of the inductive package *)
       
   103   Conseq:"[| \<forall>z::state. A |- {P' z} c {Q' z};
   103   Conseq:"[| \<forall>z::state. A |- {P' z} c {Q' z};
   104              \<forall>s t. (\<forall>z::state. P' z s --> Q' z t) --> (P s --> Q t) |] ==>
   104              \<forall>s t. (\<forall>z. P' z s --> Q' z t) --> (P s --> Q t) |] ==>
   105             A |- {P} c {Q }"
   105             A |- {P} c {Q }"
   106 
   106 
   107   (* z restricted to type state due to limitations of the inductive package *)
   107   (* z restricted to type state due to limitations of the inductive package *)
   108  eConseq:"[| \<forall>z::state. A |-e {P' z} c {Q' z};
   108  eConseq:"[| \<forall>z::state. A |-e {P' z} c {Q' z};
   109              \<forall>s v t. (\<forall>z::state. P' z s --> Q' z v t) --> (P s --> Q v t) |] ==>
   109              \<forall>s v t. (\<forall>z. P' z s --> Q' z v t) --> (P s --> Q v t) |] ==>
   110             A |-e {P} c {Q }"
   110             A |-e {P} c {Q }"
   111 
   111 
   112 
   112 
   113 subsection "Derived Rules"
   113 subsection "Derived Rules"
   114 
   114 
   145 done
   145 done
   146 
   146 
   147 lemma Union: "A |\<turnstile> (\<Union>z. C z) = (\<forall>z. A |\<turnstile> C z)"
   147 lemma Union: "A |\<turnstile> (\<Union>z. C z) = (\<forall>z. A |\<turnstile> C z)"
   148 by (auto intro: hoare_ehoare.ConjI hoare_ehoare.ConjE)
   148 by (auto intro: hoare_ehoare.ConjI hoare_ehoare.ConjE)
   149 
   149 
   150 lemma Impl': 
   150 lemma Impl1: 
   151   "\<forall>z. A\<union> (\<Union>z. (\<lambda>(C,m). (P z C m, Impl C m, Q (z::state) C m))`ms) ||- 
   151    "\<lbrakk>\<forall>z::state. A\<union> (\<Union>z. (\<lambda>M. (P z M, Impl M, Q z M))`Ms) |\<turnstile> 
   152                 (\<lambda>(C,m). (P z C m, body C m, Q (z::state) C m))`ms ==>
   152                         (\<lambda>M. (P z M, body M, Q z M))`Ms; 
   153        A    ||- (\<lambda>(C,m). (P z C m, Impl C m, Q (z::state) C m))`ms"
   153     M \<in> Ms\<rbrakk> \<Longrightarrow> 
   154 apply (drule Union[THEN iffD2])
   154                 A         \<turnstile>  {P z M} Impl M {Q z M}"
   155 apply (drule hoare_ehoare.Impl)
   155 apply (drule hoare_ehoare.Impl)
   156 apply (drule Union[THEN iffD1])
       
   157 apply (erule spec)
       
   158 done
       
   159 
       
   160 lemma Impl1: 
       
   161    "\<lbrakk>\<forall>z. A\<union> (\<Union>z. (\<lambda>(C,m). (P z C m, Impl C m, Q (z::state) C m))`ms) ||- 
       
   162                  (\<lambda>(C,m). (P z C m, body C m, Q (z::state) C m))`ms; 
       
   163     (C,m)\<in> ms\<rbrakk> \<Longrightarrow> 
       
   164          A             |- {P z C m} Impl C m {Q (z::state) C m}"
       
   165 apply (drule Impl')
       
   166 apply (erule Weaken)
   156 apply (erule Weaken)
   167 apply (auto del: image_eqI intro: rev_image_eqI)
   157 apply (auto del: image_eqI intro: rev_image_eqI)
   168 done
   158 done
   169 
   159 
   170 end
   160 end