| author | wenzelm | 
| Tue, 06 Jun 2023 11:07:49 +0200 | |
| changeset 78134 | a11ebc8c751a | 
| parent 67613 | ce654b0e6d69 | 
| child 80914 | d97fdabd9e2b | 
| 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: 
50161diff
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: 
52726diff
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: 
52726diff
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: 
52726diff
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: 
50161diff
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: | |
| 67613 | 31 | "D (dom s) c A' \<Longrightarrow> c \<noteq> SKIP \<Longrightarrow> \<exists>cs'. (c,s) \<rightarrow> cs'" | 
| 52726 | 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 | ||
| 67613 | 41 | lemma D_mono: "D A c M \<Longrightarrow> A \<subseteq> A' \<Longrightarrow> \<exists>M'. D A' c M' & M <= M'" | 
| 52726 | 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 | |
| 67406 | 48 | with If.IH \<open>A \<subseteq> A'\<close> obtain M1' M2' | 
| 52726 | 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'" | |
| 67406 | 51 | using \<open>vars b \<subseteq> A\<close> \<open>A \<subseteq> A'\<close> \<open>M = M1 \<inter> M2\<close> by(fastforce intro: D.intros)+ | 
| 52726 | 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: | |
| 67613 | 58 | "(c,s) \<rightarrow> (c',s') \<Longrightarrow> D (dom s) c A \<Longrightarrow> \<exists>A'. D (dom s') c' A' & A <= A'" | 
| 52726 | 59 | proof (induction arbitrary: A rule: small_step_induct) | 
| 60 | case (While b c s) | |
| 63540 | 61 | then obtain A' where A': "vars b \<subseteq> dom s" "A = dom s" "D (dom s) c A'" by blast | 
| 52726 | 62 | then obtain A'' where "D A' c A''" by (metis D_incr D_mono) | 
| 63540 | 63 | with A' have "D (dom s) (IF b THEN c;; WHILE b DO c ELSE SKIP) (dom s)" | 
| 67406 | 64 | by (metis D.If[OF \<open>vars b \<subseteq> dom s\<close> D.Seq[OF \<open>D (dom s) c A'\<close> D.While[OF _ \<open>D A' c A''\<close>]] D.Skip] D_incr Int_absorb1 subset_trans) | 
| 65 | thus ?case by (metis D_incr \<open>A = dom s\<close>) | |
| 52726 | 66 | next | 
| 67 | case Seq2 thus ?case by auto (metis D_mono D.intros(3)) | |
| 68 | qed (auto intro: D.intros) | |
| 69 | ||
| 70 | theorem D_sound: | |
| 71 | "(c,s) \<rightarrow>* (c',s') \<Longrightarrow> D (dom s) c A' | |
| 72 | \<Longrightarrow> (\<exists>cs''. (c',s') \<rightarrow> cs'') \<or> c' = SKIP" | |
| 73 | apply(induction arbitrary: A' rule:star_induct) | |
| 74 | apply (metis progress) | |
| 75 | by (metis D_preservation) | |
| 76 | ||
| 43158 | 77 | end |