| author | nipkow | 
| Fri, 07 Feb 2003 16:40:23 +0100 | |
| changeset 13810 | c3fbfd472365 | 
| parent 12934 | 6003b4f916c0 | 
| child 16417 | 9bc16273c2d4 | 
| permissions | -rw-r--r-- | 
| 11376 | 1  | 
(* Title: HOL/NanoJava/AxSem.thy  | 
2  | 
ID: $Id$  | 
|
3  | 
Author: David von Oheimb  | 
|
4  | 
Copyright 2001 Technische Universitaet Muenchen  | 
|
5  | 
*)  | 
|
6  | 
||
| 11560 | 7  | 
header "Axiomatic Semantics"  | 
| 11376 | 8  | 
|
9  | 
theory AxSem = State:  | 
|
10  | 
||
11  | 
types assn = "state => bool"  | 
|
| 11486 | 12  | 
vassn = "val => assn"  | 
| 11476 | 13  | 
triple = "assn \<times> stmt \<times> assn"  | 
14  | 
etriple = "assn \<times> expr \<times> vassn"  | 
|
| 11376 | 15  | 
translations  | 
16  | 
"assn" \<leftharpoondown> (type)"state => bool"  | 
|
| 11476 | 17  | 
"vassn" \<leftharpoondown> (type)"val => assn"  | 
18  | 
"triple" \<leftharpoondown> (type)"assn \<times> stmt \<times> assn"  | 
|
19  | 
"etriple" \<leftharpoondown> (type)"assn \<times> expr \<times> vassn"  | 
|
| 11376 | 20  | 
|
| 11476 | 21  | 
consts hoare :: "(triple set \<times> triple set) set"  | 
22  | 
consts ehoare :: "(triple set \<times> etriple ) set"  | 
|
| 11376 | 23  | 
syntax (xsymbols)  | 
24  | 
 "@hoare"  :: "[triple set,  triple set    ] => bool" ("_ |\<turnstile>/ _" [61,61]    60)
 | 
|
25  | 
"@hoare1" :: "[triple set, assn,stmt,assn] => bool"  | 
|
26  | 
                                   ("_ \<turnstile>/ ({(1_)}/ (_)/ {(1_)})" [61,3,90,3]60)
 | 
|
| 11486 | 27  | 
"@ehoare"  :: "[triple set,  etriple       ] => bool" ("_ |\<turnstile>\<^sub>e/ _"[61,61]60)
 | 
| 11476 | 28  | 
"@ehoare1" :: "[triple set, assn,expr,vassn]=> bool"  | 
| 11486 | 29  | 
                                  ("_ \<turnstile>\<^sub>e/ ({(1_)}/ (_)/ {(1_)})" [61,3,90,3]60)
 | 
| 11376 | 30  | 
syntax  | 
31  | 
 "@hoare"  :: "[triple set,  triple set    ] => bool" ("_ ||-/ _" [61,61] 60)
 | 
|
32  | 
"@hoare1" :: "[triple set, assn,stmt,assn] => bool"  | 
|
33  | 
                                  ("_ |-/ ({(1_)}/ (_)/ {(1_)})" [61,3,90,3] 60)
 | 
|
| 11476 | 34  | 
"@ehoare"  :: "[triple set,  etriple       ] => bool" ("_ ||-e/ _"[61,61] 60)
 | 
35  | 
"@ehoare1" :: "[triple set, assn,expr,vassn]=> bool"  | 
|
36  | 
                                 ("_ |-e/ ({(1_)}/ (_)/ {(1_)})" [61,3,90,3] 60)
 | 
|
| 11376 | 37  | 
|
| 11476 | 38  | 
translations "A |\<turnstile> C" \<rightleftharpoons> "(A,C) \<in> hoare"  | 
39  | 
             "A  \<turnstile> {P}c{Q}"  \<rightleftharpoons> "A |\<turnstile> {(P,c,Q)}"
 | 
|
| 11486 | 40  | 
"A |\<turnstile>\<^sub>e t" \<rightleftharpoons> "(A,t) \<in> ehoare"  | 
| 
11558
 
6539627881e8
simplified vnam/vname, introduced fname, improved comments
 
oheimb 
parents: 
11508 
diff
changeset
 | 
41  | 
"A |\<turnstile>\<^sub>e (P,e,Q)" \<rightleftharpoons> "(A,P,e,Q) \<in> ehoare" (* shouldn't be necessary *)  | 
| 11486 | 42  | 
             "A  \<turnstile>\<^sub>e {P}e{Q}" \<rightleftharpoons> "A |\<turnstile>\<^sub>e (P,e,Q)"
 | 
| 11376 | 43  | 
|
| 11476 | 44  | 
|
| 11560 | 45  | 
subsection "Hoare Logic Rules"  | 
46  | 
||
| 11476 | 47  | 
inductive hoare ehoare  | 
| 11376 | 48  | 
intros  | 
49  | 
||
50  | 
  Skip:  "A |- {P} Skip {P}"
 | 
|
51  | 
||
52  | 
  Comp: "[| A |- {P} c1 {Q}; A |- {Q} c2 {R} |] ==> A |- {P} c1;;c2 {R}"
 | 
|
53  | 
||
| 11476 | 54  | 
  Cond: "[| A |-e {P} e {Q}; 
 | 
55  | 
            \<forall>v. A |- {Q v} (if v \<noteq> Null then c1 else c2) {R} |] ==>
 | 
|
56  | 
            A |- {P} If(e) c1 Else c2 {R}"
 | 
|
57  | 
||
58  | 
  Loop: "A |- {\<lambda>s. P s \<and> s<x> \<noteq> Null} c {P} ==>
 | 
|
59  | 
         A |- {P} While(x) c {\<lambda>s. P s \<and> s<x> = Null}"
 | 
|
| 11376 | 60  | 
|
| 11476 | 61  | 
  LAcc: "A |-e {\<lambda>s. P (s<x>) s} LAcc x {P}"
 | 
| 11376 | 62  | 
|
| 11476 | 63  | 
  LAss: "A |-e {P} e {\<lambda>v s.  Q (lupd(x\<mapsto>v) s)} ==>
 | 
64  | 
         A |-  {P} x:==e {Q}"
 | 
|
65  | 
||
66  | 
  FAcc: "A |-e {P} e {\<lambda>v s. \<forall>a. v=Addr a --> Q (get_field s a f) s} ==>
 | 
|
67  | 
         A |-e {P} e..f {Q}"
 | 
|
| 11376 | 68  | 
|
| 11476 | 69  | 
  FAss: "[| A |-e {P} e1 {\<lambda>v s. \<forall>a. v=Addr a --> Q a s};
 | 
70  | 
        \<forall>a. A |-e {Q a} e2 {\<lambda>v s. R (upd_obj a f v s)} |] ==>
 | 
|
71  | 
            A |-  {P} e1..f:==e2 {R}"
 | 
|
| 11376 | 72  | 
|
| 11476 | 73  | 
  NewC: "A |-e {\<lambda>s. \<forall>a. new_Addr s = Addr a --> P (Addr a) (new_obj a C s)}
 | 
74  | 
                new C {P}"
 | 
|
| 11376 | 75  | 
|
| 11476 | 76  | 
  Cast: "A |-e {P} e {\<lambda>v s. (case v of Null => True 
 | 
77  | 
| Addr a => obj_class s a <=C C) --> Q v s} ==>  | 
|
78  | 
         A |-e {P} Cast C e {Q}"
 | 
|
| 11376 | 79  | 
|
| 11476 | 80  | 
  Call: "[| A |-e {P} e1 {Q}; \<forall>a. A |-e {Q a} e2 {R a};
 | 
| 11565 | 81  | 
    \<forall>a p ls. A |- {\<lambda>s'. \<exists>s. R a p s \<and> ls = s \<and> 
 | 
| 11772 | 82  | 
s' = lupd(This\<mapsto>a)(lupd(Par\<mapsto>p)(del_locs s))}  | 
| 11565 | 83  | 
                  Meth (C,m) {\<lambda>s. S (s<Res>) (set_locs ls s)} |] ==>
 | 
| 11476 | 84  | 
             A |-e {P} {C}e1..m(e2) {S}"
 | 
| 11376 | 85  | 
|
| 
11497
 
0e66e0114d9a
corrected initialization of locals, streamlined Impl
 
oheimb 
parents: 
11486 
diff
changeset
 | 
86  | 
  Meth: "\<forall>D. A |- {\<lambda>s'. \<exists>s a. s<This> = Addr a \<and> D = obj_class s a \<and> D <=C C \<and> 
 | 
| 
 
0e66e0114d9a
corrected initialization of locals, streamlined Impl
 
oheimb 
parents: 
11486 
diff
changeset
 | 
87  | 
P s \<and> s' = init_locs D m s}  | 
| 
 
0e66e0114d9a
corrected initialization of locals, streamlined Impl
 
oheimb 
parents: 
11486 
diff
changeset
 | 
88  | 
                  Impl (D,m) {Q} ==>
 | 
| 11507 | 89  | 
             A |- {P} Meth (C,m) {Q}"
 | 
| 11376 | 90  | 
|
| 11565 | 91  | 
  --{* @{text "\<Union>Z"} instead of @{text "\<forall>Z"} in the conclusion and\\
 | 
92  | 
Z restricted to type state due to limitations of the inductive package *}  | 
|
93  | 
Impl: "\<forall>Z::state. A\<union> (\<Union>Z. (\<lambda>Cm. (P Z Cm, Impl Cm, Q Z Cm))`Ms) ||-  | 
|
94  | 
(\<lambda>Cm. (P Z Cm, body Cm, Q Z Cm))`Ms ==>  | 
|
95  | 
A ||- (\<lambda>Cm. (P Z Cm, Impl Cm, Q Z Cm))`Ms"  | 
|
| 11376 | 96  | 
|
| 
11558
 
6539627881e8
simplified vnam/vname, introduced fname, improved comments
 
oheimb 
parents: 
11508 
diff
changeset
 | 
97  | 
--{* structural rules *}
 | 
| 11376 | 98  | 
|
| 11476 | 99  | 
  Asm:  "   a \<in> A ==> A ||- {a}"
 | 
100  | 
||
101  | 
  ConjI: " \<forall>c \<in> C. A ||- {c} ==> A ||- C"
 | 
|
102  | 
||
103  | 
  ConjE: "[|A ||- C; c \<in> C |] ==> A ||- {c}"
 | 
|
104  | 
||
| 11565 | 105  | 
  --{* Z restricted to type state due to limitations of the inductive package *}
 | 
106  | 
  Conseq:"[| \<forall>Z::state. A |- {P' Z} c {Q' Z};
 | 
|
107  | 
\<forall>s t. (\<forall>Z. P' Z s --> Q' Z t) --> (P s --> Q t) |] ==>  | 
|
| 11376 | 108  | 
            A |- {P} c {Q }"
 | 
109  | 
||
| 11565 | 110  | 
  --{* Z restricted to type state due to limitations of the inductive package *}
 | 
111  | 
 eConseq:"[| \<forall>Z::state. A |-e {P' Z} e {Q' Z};
 | 
|
112  | 
\<forall>s v t. (\<forall>Z. P' Z s --> Q' Z v t) --> (P s --> Q v t) |] ==>  | 
|
113  | 
            A |-e {P} e {Q }"
 | 
|
114  | 
||
115  | 
||
| 
12934
 
6003b4f916c0
Clarification wrt. use of polymorphic variants of Hoare logic rules
 
oheimb 
parents: 
11772 
diff
changeset
 | 
116  | 
subsection "Fully polymorphic variants, required for Example only"  | 
| 11376 | 117  | 
|
| 11565 | 118  | 
axioms  | 
| 
12934
 
6003b4f916c0
Clarification wrt. use of polymorphic variants of Hoare logic rules
 
oheimb 
parents: 
11772 
diff
changeset
 | 
119  | 
|
| 11565 | 120  | 
  Conseq:"[| \<forall>Z. A |- {P' Z} c {Q' Z};
 | 
121  | 
\<forall>s t. (\<forall>Z. P' Z s --> Q' Z t) --> (P s --> Q t) |] ==>  | 
|
122  | 
                 A |- {P} c {Q }"
 | 
|
123  | 
||
124  | 
 eConseq:"[| \<forall>Z. A |-e {P' Z} e {Q' Z};
 | 
|
125  | 
\<forall>s v t. (\<forall>Z. P' Z s --> Q' Z v t) --> (P s --> Q v t) |] ==>  | 
|
126  | 
                 A |-e {P} e {Q }"
 | 
|
127  | 
||
128  | 
Impl: "\<forall>Z. A\<union> (\<Union>Z. (\<lambda>Cm. (P Z Cm, Impl Cm, Q Z Cm))`Ms) ||-  | 
|
129  | 
(\<lambda>Cm. (P Z Cm, body Cm, Q Z Cm))`Ms ==>  | 
|
130  | 
A ||- (\<lambda>Cm. (P Z Cm, Impl Cm, Q Z Cm))`Ms"  | 
|
| 11376 | 131  | 
|
132  | 
subsection "Derived Rules"  | 
|
133  | 
||
134  | 
lemma Conseq1: "\<lbrakk>A \<turnstile> {P'} c {Q}; \<forall>s. P s \<longrightarrow> P' s\<rbrakk> \<Longrightarrow> A \<turnstile> {P} c {Q}"
 | 
|
| 11476 | 135  | 
apply (rule hoare_ehoare.Conseq)  | 
136  | 
apply (rule allI, assumption)  | 
|
137  | 
apply fast  | 
|
138  | 
done  | 
|
139  | 
||
140  | 
lemma Conseq2: "\<lbrakk>A \<turnstile> {P} c {Q'}; \<forall>t. Q' t \<longrightarrow> Q t\<rbrakk> \<Longrightarrow> A \<turnstile> {P} c {Q}"
 | 
|
141  | 
apply (rule hoare_ehoare.Conseq)  | 
|
142  | 
apply (rule allI, assumption)  | 
|
143  | 
apply fast  | 
|
144  | 
done  | 
|
145  | 
||
| 11486 | 146  | 
lemma eConseq1: "\<lbrakk>A \<turnstile>\<^sub>e {P'} e {Q}; \<forall>s. P s \<longrightarrow> P' s\<rbrakk> \<Longrightarrow> A \<turnstile>\<^sub>e {P} e {Q}"
 | 
| 11476 | 147  | 
apply (rule hoare_ehoare.eConseq)  | 
148  | 
apply (rule allI, assumption)  | 
|
149  | 
apply fast  | 
|
150  | 
done  | 
|
151  | 
||
| 11486 | 152  | 
lemma eConseq2: "\<lbrakk>A \<turnstile>\<^sub>e {P} e {Q'}; \<forall>v t. Q' v t \<longrightarrow> Q v t\<rbrakk> \<Longrightarrow> A \<turnstile>\<^sub>e {P} e {Q}"
 | 
| 11476 | 153  | 
apply (rule hoare_ehoare.eConseq)  | 
| 11376 | 154  | 
apply (rule allI, assumption)  | 
155  | 
apply fast  | 
|
156  | 
done  | 
|
157  | 
||
158  | 
lemma Weaken: "\<lbrakk>A |\<turnstile> C'; C \<subseteq> C'\<rbrakk> \<Longrightarrow> A |\<turnstile> C"  | 
|
| 11476 | 159  | 
apply (rule hoare_ehoare.ConjI)  | 
| 11376 | 160  | 
apply clarify  | 
| 11476 | 161  | 
apply (drule hoare_ehoare.ConjE)  | 
| 11376 | 162  | 
apply fast  | 
163  | 
apply assumption  | 
|
164  | 
done  | 
|
165  | 
||
| 11565 | 166  | 
lemma Thin_lemma:  | 
167  | 
"(A' |\<turnstile> C \<longrightarrow> (\<forall>A. A' \<subseteq> A \<longrightarrow> A |\<turnstile> C )) \<and>  | 
|
168  | 
   (A'  \<turnstile>\<^sub>e {P} e {Q} \<longrightarrow> (\<forall>A. A' \<subseteq> A \<longrightarrow> A  \<turnstile>\<^sub>e {P} e {Q}))"
 | 
|
169  | 
apply (rule hoare_ehoare.induct)  | 
|
170  | 
apply (tactic "ALLGOALS(EVERY'[Clarify_tac, REPEAT o smp_tac 1])")  | 
|
171  | 
apply (blast intro: hoare_ehoare.Skip)  | 
|
172  | 
apply (blast intro: hoare_ehoare.Comp)  | 
|
173  | 
apply (blast intro: hoare_ehoare.Cond)  | 
|
174  | 
apply (blast intro: hoare_ehoare.Loop)  | 
|
175  | 
apply (blast intro: hoare_ehoare.LAcc)  | 
|
176  | 
apply (blast intro: hoare_ehoare.LAss)  | 
|
177  | 
apply (blast intro: hoare_ehoare.FAcc)  | 
|
178  | 
apply (blast intro: hoare_ehoare.FAss)  | 
|
179  | 
apply (blast intro: hoare_ehoare.NewC)  | 
|
180  | 
apply (blast intro: hoare_ehoare.Cast)  | 
|
181  | 
apply (erule hoare_ehoare.Call)  | 
|
182  | 
apply (rule, drule spec, erule conjE, tactic "smp_tac 1 1", assumption)  | 
|
183  | 
apply blast  | 
|
184  | 
apply (blast intro!: hoare_ehoare.Meth)  | 
|
185  | 
apply (blast intro!: hoare_ehoare.Impl)  | 
|
186  | 
apply (blast intro!: hoare_ehoare.Asm)  | 
|
187  | 
apply (blast intro: hoare_ehoare.ConjI)  | 
|
188  | 
apply (blast intro: hoare_ehoare.ConjE)  | 
|
189  | 
apply (rule hoare_ehoare.Conseq)  | 
|
190  | 
apply (rule, drule spec, erule conjE, tactic "smp_tac 1 1", assumption+)  | 
|
191  | 
apply (rule hoare_ehoare.eConseq)  | 
|
192  | 
apply (rule, drule spec, erule conjE, tactic "smp_tac 1 1", assumption+)  | 
|
193  | 
done  | 
|
194  | 
||
195  | 
lemma cThin: "\<lbrakk>A' |\<turnstile> C; A' \<subseteq> A\<rbrakk> \<Longrightarrow> A |\<turnstile> C"  | 
|
196  | 
by (erule (1) conjunct1 [OF Thin_lemma, rule_format])  | 
|
197  | 
||
198  | 
lemma eThin: "\<lbrakk>A' \<turnstile>\<^sub>e {P} e {Q}; A' \<subseteq> A\<rbrakk> \<Longrightarrow> A \<turnstile>\<^sub>e {P} e {Q}"
 | 
|
199  | 
by (erule (1) conjunct2 [OF Thin_lemma, rule_format])  | 
|
200  | 
||
201  | 
||
202  | 
lemma Union: "A |\<turnstile> (\<Union>Z. C Z) = (\<forall>Z. A |\<turnstile> C Z)"  | 
|
| 11476 | 203  | 
by (auto intro: hoare_ehoare.ConjI hoare_ehoare.ConjE)  | 
| 11376 | 204  | 
|
| 11565 | 205  | 
lemma Impl1':  | 
| 
12934
 
6003b4f916c0
Clarification wrt. use of polymorphic variants of Hoare logic rules
 
oheimb 
parents: 
11772 
diff
changeset
 | 
206  | 
"\<lbrakk>\<forall>Z::state. A\<union> (\<Union>Z. (\<lambda>Cm. (P Z Cm, Impl Cm, Q Z Cm))`Ms) |\<turnstile>  | 
| 11565 | 207  | 
(\<lambda>Cm. (P Z Cm, body Cm, Q Z Cm))`Ms;  | 
| 11508 | 208  | 
Cm \<in> Ms\<rbrakk> \<Longrightarrow>  | 
| 11565 | 209  | 
                A   \<turnstile>  {P Z Cm} Impl Cm {Q Z Cm}"
 | 
| 
12934
 
6003b4f916c0
Clarification wrt. use of polymorphic variants of Hoare logic rules
 
oheimb 
parents: 
11772 
diff
changeset
 | 
210  | 
apply (drule AxSem.Impl)  | 
| 11376 | 211  | 
apply (erule Weaken)  | 
212  | 
apply (auto del: image_eqI intro: rev_image_eqI)  | 
|
213  | 
done  | 
|
214  | 
||
| 11565 | 215  | 
lemmas Impl1 = AxSem.Impl [of _ _ _ "{Cm}", simplified, standard]
 | 
216  | 
||
| 11376 | 217  | 
end  |