| author | wenzelm | 
| Sun, 24 Feb 2013 14:14:07 +0100 | |
| changeset 51267 | c68c1b89a0f1 | 
| parent 50161 | 4fc4237488ab | 
| child 51466 | d53cdbca1be4 | 
| permissions | -rw-r--r-- | 
| 43158 | 1 | (* Author: Tobias Nipkow *) | 
| 2 | ||
| 50161 | 3 | theory Def_Init_Sound_Small | 
| 4 | imports Def_Init Def_Init_Small | |
| 43158 | 5 | begin | 
| 6 | ||
| 7 | subsection "Soundness wrt Small Steps" | |
| 8 | ||
| 9 | theorem progress: | |
| 10 | "D (dom s) c A' \<Longrightarrow> c \<noteq> SKIP \<Longrightarrow> EX cs'. (c,s) \<rightarrow> cs'" | |
| 45015 | 11 | proof (induction c arbitrary: s A') | 
| 43158 | 12 | case Assign thus ?case by auto (metis aval_Some small_step.Assign) | 
| 13 | next | |
| 14 | case (If b c1 c2) | |
| 15 | then obtain bv where "bval b s = Some bv" by (auto dest!:bval_Some) | |
| 16 | then show ?case | |
| 17 | by(cases bv)(auto intro: small_step.IfTrue small_step.IfFalse) | |
| 44890 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 nipkow parents: 
43158diff
changeset | 18 | qed (fastforce intro: small_step.intros)+ | 
| 43158 | 19 | |
| 20 | lemma D_mono: "D A c M \<Longrightarrow> A \<subseteq> A' \<Longrightarrow> EX M'. D A' c M' & M <= M'" | |
| 45015 | 21 | proof (induction c arbitrary: A A' M) | 
| 47818 | 22 | case Seq thus ?case by auto (metis D.intros(3)) | 
| 43158 | 23 | next | 
| 24 | case (If b c1 c2) | |
| 25 | then obtain M1 M2 where "vars b \<subseteq> A" "D A c1 M1" "D A c2 M2" "M = M1 \<inter> M2" | |
| 26 | by auto | |
| 45015 | 27 | with If.IH `A \<subseteq> A'` obtain M1' M2' | 
| 43158 | 28 | where "D A' c1 M1'" "D A' c2 M2'" and "M1 \<subseteq> M1'" "M2 \<subseteq> M2'" by metis | 
| 29 | hence "D A' (IF b THEN c1 ELSE c2) (M1' \<inter> M2')" and "M \<subseteq> M1' \<inter> M2'" | |
| 44890 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 nipkow parents: 
43158diff
changeset | 30 | using `vars b \<subseteq> A` `A \<subseteq> A'` `M = M1 \<inter> M2` by(fastforce intro: D.intros)+ | 
| 43158 | 31 | thus ?case by metis | 
| 32 | next | |
| 33 | case While thus ?case by auto (metis D.intros(5) subset_trans) | |
| 34 | qed (auto intro: D.intros) | |
| 35 | ||
| 36 | theorem D_preservation: | |
| 37 | "(c,s) \<rightarrow> (c',s') \<Longrightarrow> D (dom s) c A \<Longrightarrow> EX A'. D (dom s') c' A' & A <= A'" | |
| 45015 | 38 | proof (induction arbitrary: A rule: small_step_induct) | 
| 43158 | 39 | case (While b c s) | 
| 40 | then obtain A' where "vars b \<subseteq> dom s" "A = dom s" "D (dom s) c A'" by blast | |
| 41 | moreover | |
| 42 | then obtain A'' where "D A' c A''" by (metis D_incr D_mono) | |
| 43 | ultimately have "D (dom s) (IF b THEN c; WHILE b DO c ELSE SKIP) (dom s)" | |
| 47818 | 44 | 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) | 
| 43158 | 45 | thus ?case by (metis D_incr `A = dom s`) | 
| 46 | next | |
| 47818 | 47 | case Seq2 thus ?case by auto (metis D_mono D.intros(3)) | 
| 43158 | 48 | qed (auto intro: D.intros) | 
| 49 | ||
| 50 | theorem D_sound: | |
| 51 | "(c,s) \<rightarrow>* (c',s') \<Longrightarrow> D (dom s) c A' \<Longrightarrow> c' \<noteq> SKIP | |
| 52 | \<Longrightarrow> \<exists>cs''. (c',s') \<rightarrow> cs''" | |
| 45015 | 53 | apply(induction arbitrary: A' rule:star_induct) | 
| 43158 | 54 | apply (metis progress) | 
| 55 | by (metis D_preservation) | |
| 56 | ||
| 57 | end |