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