| author | wenzelm | 
| Mon, 10 Feb 2014 22:08:18 +0100 | |
| changeset 55385 | 169e12bbf9a3 | 
| parent 53015 | a1119cf551e8 | 
| child 63540 | f8652d0534fa | 
| permissions | -rw-r--r-- | 
| 43158 | 1  | 
(* Author: Tobias Nipkow *)  | 
2  | 
||
| 50161 | 3  | 
theory Def_Init_Small  | 
| 52726 | 4  | 
imports Star Def_Init_Exp Def_Init  | 
| 43158 | 5  | 
begin  | 
6  | 
||
7  | 
subsection "Initialization-Sensitive Small Step Semantics"  | 
|
8  | 
||
9  | 
inductive  | 
|
10  | 
small_step :: "(com \<times> state) \<Rightarrow> (com \<times> state) \<Rightarrow> bool" (infix "\<rightarrow>" 55)  | 
|
11  | 
where  | 
|
12  | 
Assign: "aval a s = Some i \<Longrightarrow> (x ::= a, s) \<rightarrow> (SKIP, s(x := Some i))" |  | 
|
13  | 
||
| 
52046
 
bc01725d7918
replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
 
nipkow 
parents: 
50161 
diff
changeset
 | 
14  | 
Seq1: "(SKIP;;c,s) \<rightarrow> (c,s)" |  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52726 
diff
changeset
 | 
15  | 
Seq2: "(c\<^sub>1,s) \<rightarrow> (c\<^sub>1',s') \<Longrightarrow> (c\<^sub>1;;c\<^sub>2,s) \<rightarrow> (c\<^sub>1';;c\<^sub>2,s')" |  | 
| 43158 | 16  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52726 
diff
changeset
 | 
17  | 
IfTrue: "bval b s = Some True \<Longrightarrow> (IF b THEN c\<^sub>1 ELSE c\<^sub>2,s) \<rightarrow> (c\<^sub>1,s)" |  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52726 
diff
changeset
 | 
18  | 
IfFalse: "bval b s = Some False \<Longrightarrow> (IF b THEN c\<^sub>1 ELSE c\<^sub>2,s) \<rightarrow> (c\<^sub>2,s)" |  | 
| 43158 | 19  | 
|
| 
52046
 
bc01725d7918
replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
 
nipkow 
parents: 
50161 
diff
changeset
 | 
20  | 
While: "(WHILE b DO c,s) \<rightarrow> (IF b THEN c;; WHILE b DO c ELSE SKIP,s)"  | 
| 43158 | 21  | 
|
22  | 
lemmas small_step_induct = small_step.induct[split_format(complete)]  | 
|
23  | 
||
24  | 
abbreviation small_steps :: "com * state \<Rightarrow> com * state \<Rightarrow> bool" (infix "\<rightarrow>*" 55)  | 
|
25  | 
where "x \<rightarrow>* y == star small_step x y"  | 
|
26  | 
||
| 52726 | 27  | 
|
28  | 
subsection "Soundness wrt Small Steps"  | 
|
29  | 
||
30  | 
theorem progress:  | 
|
31  | 
"D (dom s) c A' \<Longrightarrow> c \<noteq> SKIP \<Longrightarrow> EX cs'. (c,s) \<rightarrow> cs'"  | 
|
32  | 
proof (induction c arbitrary: s A')  | 
|
33  | 
case Assign thus ?case by auto (metis aval_Some small_step.Assign)  | 
|
34  | 
next  | 
|
35  | 
case (If b c1 c2)  | 
|
36  | 
then obtain bv where "bval b s = Some bv" by (auto dest!:bval_Some)  | 
|
37  | 
then show ?case  | 
|
38  | 
by(cases bv)(auto intro: small_step.IfTrue small_step.IfFalse)  | 
|
39  | 
qed (fastforce intro: small_step.intros)+  | 
|
40  | 
||
41  | 
lemma D_mono: "D A c M \<Longrightarrow> A \<subseteq> A' \<Longrightarrow> EX M'. D A' c M' & M <= M'"  | 
|
42  | 
proof (induction c arbitrary: A A' M)  | 
|
43  | 
case Seq thus ?case by auto (metis D.intros(3))  | 
|
44  | 
next  | 
|
45  | 
case (If b c1 c2)  | 
|
46  | 
then obtain M1 M2 where "vars b \<subseteq> A" "D A c1 M1" "D A c2 M2" "M = M1 \<inter> M2"  | 
|
47  | 
by auto  | 
|
48  | 
with If.IH `A \<subseteq> A'` obtain M1' M2'  | 
|
49  | 
where "D A' c1 M1'" "D A' c2 M2'" and "M1 \<subseteq> M1'" "M2 \<subseteq> M2'" by metis  | 
|
50  | 
hence "D A' (IF b THEN c1 ELSE c2) (M1' \<inter> M2')" and "M \<subseteq> M1' \<inter> M2'"  | 
|
51  | 
using `vars b \<subseteq> A` `A \<subseteq> A'` `M = M1 \<inter> M2` by(fastforce intro: D.intros)+  | 
|
52  | 
thus ?case by metis  | 
|
53  | 
next  | 
|
54  | 
case While thus ?case by auto (metis D.intros(5) subset_trans)  | 
|
55  | 
qed (auto intro: D.intros)  | 
|
56  | 
||
57  | 
theorem D_preservation:  | 
|
58  | 
"(c,s) \<rightarrow> (c',s') \<Longrightarrow> D (dom s) c A \<Longrightarrow> EX A'. D (dom s') c' A' & A <= A'"  | 
|
59  | 
proof (induction arbitrary: A rule: small_step_induct)  | 
|
60  | 
case (While b c s)  | 
|
61  | 
then obtain A' where "vars b \<subseteq> dom s" "A = dom s" "D (dom s) c A'" by blast  | 
|
62  | 
moreover  | 
|
63  | 
then obtain A'' where "D A' c A''" by (metis D_incr D_mono)  | 
|
64  | 
ultimately have "D (dom s) (IF b THEN c;; WHILE b DO c ELSE SKIP) (dom s)"  | 
|
65  | 
by (metis D.If[OF `vars b \<subseteq> dom s` D.Seq[OF `D (dom s) c A'` D.While[OF _ `D A' c A''`]] D.Skip] D_incr Int_absorb1 subset_trans)  | 
|
66  | 
thus ?case by (metis D_incr `A = dom s`)  | 
|
67  | 
next  | 
|
68  | 
case Seq2 thus ?case by auto (metis D_mono D.intros(3))  | 
|
69  | 
qed (auto intro: D.intros)  | 
|
70  | 
||
71  | 
theorem D_sound:  | 
|
72  | 
"(c,s) \<rightarrow>* (c',s') \<Longrightarrow> D (dom s) c A'  | 
|
73  | 
\<Longrightarrow> (\<exists>cs''. (c',s') \<rightarrow> cs'') \<or> c' = SKIP"  | 
|
74  | 
apply(induction arbitrary: A' rule:star_induct)  | 
|
75  | 
apply (metis progress)  | 
|
76  | 
by (metis D_preservation)  | 
|
77  | 
||
| 43158 | 78  | 
end  |