9 sup (infixl "\<squnion>" 65) and |
9 sup (infixl "\<squnion>" 65) and |
10 inf (infixl "\<sqinter>" 70) and |
10 inf (infixl "\<sqinter>" 70) and |
11 bot ("\<bottom>") and |
11 bot ("\<bottom>") and |
12 top ("\<top>") |
12 top ("\<top>") |
13 |
13 |
14 fun Step :: "(vname \<Rightarrow> aexp \<Rightarrow> 'a \<Rightarrow> 'a::sup) \<Rightarrow> (bexp \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a acom \<Rightarrow> 'a acom" where |
14 context |
15 "Step f g S (SKIP {Q}) = (SKIP {S})" | |
15 fixes f :: "vname \<Rightarrow> aexp \<Rightarrow> 'a \<Rightarrow> 'a::sup" |
16 "Step f g S (x ::= e {Q}) = |
16 fixes g :: "bexp \<Rightarrow> 'a \<Rightarrow> 'a" |
|
17 begin |
|
18 fun Step :: "'a \<Rightarrow> 'a acom \<Rightarrow> 'a acom" where |
|
19 "Step S (SKIP {Q}) = (SKIP {S})" | |
|
20 "Step S (x ::= e {Q}) = |
17 x ::= e {f x e S}" | |
21 x ::= e {f x e S}" | |
18 "Step f g S (C1; C2) = Step f g S C1; Step f g (post C1) C2" | |
22 "Step S (C1; C2) = Step S C1; Step (post C1) C2" | |
19 "Step f g S (IF b THEN {P1} C1 ELSE {P2} C2 {Q}) = |
23 "Step S (IF b THEN {P1} C1 ELSE {P2} C2 {Q}) = |
20 IF b THEN {g b S} Step f g P1 C1 ELSE {g (Not b) S} Step f g P2 C2 |
24 IF b THEN {g b S} Step P1 C1 ELSE {g (Not b) S} Step P2 C2 |
21 {post C1 \<squnion> post C2}" | |
25 {post C1 \<squnion> post C2}" | |
22 "Step f g S ({I} WHILE b DO {P} C {Q}) = |
26 "Step S ({I} WHILE b DO {P} C {Q}) = |
23 {S \<squnion> post C} WHILE b DO {g b I} Step f g P C {g (Not b) I}" |
27 {S \<squnion> post C} WHILE b DO {g b I} Step P C {g (Not b) I}" |
|
28 end |
24 |
29 |
25 lemma strip_Step[simp]: "strip(Step f g S C) = strip C" |
30 lemma strip_Step[simp]: "strip(Step f g S C) = strip C" |
26 by(induct C arbitrary: S) auto |
31 by(induct C arbitrary: S) auto |
27 |
32 |
28 |
33 |
31 subsubsection "Annotated commands as a complete lattice" |
36 subsubsection "Annotated commands as a complete lattice" |
32 |
37 |
33 instantiation acom :: (order) order |
38 instantiation acom :: (order) order |
34 begin |
39 begin |
35 |
40 |
36 fun less_eq_acom :: "('a::order)acom \<Rightarrow> 'a acom \<Rightarrow> bool" where |
41 definition less_eq_acom :: "('a::order)acom \<Rightarrow> 'a acom \<Rightarrow> bool" where |
37 "(SKIP {P}) \<le> (SKIP {P'}) = (P \<le> P')" | |
42 "C1 \<le> C2 \<longleftrightarrow> strip C1 = strip C2 \<and> (\<forall>p<size(annos C1). anno C1 p \<le> anno C2 p)" |
38 "(x ::= e {P}) \<le> (x' ::= e' {P'}) = (x=x' \<and> e=e' \<and> P \<le> P')" | |
|
39 "(C1;C2) \<le> (C1';C2') = (C1 \<le> C1' \<and> C2 \<le> C2')" | |
|
40 "(IF b THEN {P1} C1 ELSE {P2} C2 {Q}) \<le> (IF b' THEN {P1'} C1' ELSE {P2'} C2' {Q'}) = |
|
41 (b=b' \<and> P1 \<le> P1' \<and> C1 \<le> C1' \<and> P2 \<le> P2' \<and> C2 \<le> C2' \<and> Q \<le> Q')" | |
|
42 "({I} WHILE b DO {P} C {Q}) \<le> ({I'} WHILE b' DO {P'} C' {Q'}) = |
|
43 (b=b' \<and> C \<le> C' \<and> I \<le> I' \<and> P \<le> P' \<and> Q \<le> Q')" | |
|
44 "less_eq_acom _ _ = False" |
|
45 |
|
46 lemma SKIP_le: "SKIP {S} \<le> c \<longleftrightarrow> (\<exists>S'. c = SKIP {S'} \<and> S \<le> S')" |
|
47 by (cases c) auto |
|
48 |
|
49 lemma Assign_le: "x ::= e {S} \<le> c \<longleftrightarrow> (\<exists>S'. c = x ::= e {S'} \<and> S \<le> S')" |
|
50 by (cases c) auto |
|
51 |
|
52 lemma Seq_le: "C1;C2 \<le> C \<longleftrightarrow> (\<exists>C1' C2'. C = C1';C2' \<and> C1 \<le> C1' \<and> C2 \<le> C2')" |
|
53 by (cases C) auto |
|
54 |
|
55 lemma If_le: "IF b THEN {p1} C1 ELSE {p2} C2 {S} \<le> C \<longleftrightarrow> |
|
56 (\<exists>p1' p2' C1' C2' S'. C = IF b THEN {p1'} C1' ELSE {p2'} C2' {S'} \<and> |
|
57 p1 \<le> p1' \<and> p2 \<le> p2' \<and> C1 \<le> C1' \<and> C2 \<le> C2' \<and> S \<le> S')" |
|
58 by (cases C) auto |
|
59 |
|
60 lemma While_le: "{I} WHILE b DO {p} C {P} \<le> W \<longleftrightarrow> |
|
61 (\<exists>I' p' C' P'. W = {I'} WHILE b DO {p'} C' {P'} \<and> C \<le> C' \<and> p \<le> p' \<and> I \<le> I' \<and> P \<le> P')" |
|
62 by (cases W) auto |
|
63 |
43 |
64 definition less_acom :: "'a acom \<Rightarrow> 'a acom \<Rightarrow> bool" where |
44 definition less_acom :: "'a acom \<Rightarrow> 'a acom \<Rightarrow> bool" where |
65 "less_acom x y = (x \<le> y \<and> \<not> y \<le> x)" |
45 "less_acom x y = (x \<le> y \<and> \<not> y \<le> x)" |
66 |
46 |
67 instance |
47 instance |
68 proof |
48 proof |
69 case goal1 show ?case by(simp add: less_acom_def) |
49 case goal1 show ?case by(simp add: less_acom_def) |
70 next |
50 next |
71 case goal2 thus ?case by (induct x) auto |
51 case goal2 thus ?case by(auto simp: less_eq_acom_def) |
72 next |
52 next |
73 case goal3 thus ?case |
53 case goal3 thus ?case by(fastforce simp: less_eq_acom_def size_annos) |
74 apply(induct x y arbitrary: z rule: less_eq_acom.induct) |
|
75 apply (auto intro: le_trans simp: SKIP_le Assign_le Seq_le If_le While_le) |
|
76 done |
|
77 next |
54 next |
78 case goal4 thus ?case |
55 case goal4 thus ?case |
79 apply(induct x y rule: less_eq_acom.induct) |
56 by(fastforce simp: le_antisym less_eq_acom_def size_annos |
80 apply (auto intro: le_antisym) |
57 eq_acom_iff_strip_anno) |
81 done |
|
82 qed |
58 qed |
83 |
59 |
84 end |
60 end |
85 |
61 |
86 text_raw{*\snip{subadef}{2}{1}{% *} |
62 lemma less_eq_acom_annos: |
87 fun sub\<^isub>1 :: "'a acom \<Rightarrow> 'a acom" where |
63 "C1 \<le> C2 \<longleftrightarrow> strip C1 = strip C2 \<and> list_all2 (op \<le>) (annos C1) (annos C2)" |
88 "sub\<^isub>1(C\<^isub>1;C\<^isub>2) = C\<^isub>1" | |
64 by(auto simp add: less_eq_acom_def anno_def list_all2_conv_all_nth size_annos_same2) |
89 "sub\<^isub>1(IF b THEN {P\<^isub>1} C\<^isub>1 ELSE {P\<^isub>2} C\<^isub>2 {Q}) = C\<^isub>1" | |
65 |
90 "sub\<^isub>1({I} WHILE b DO {P} C {Q}) = C" |
66 lemma SKIP_le[simp]: "SKIP {S} \<le> c \<longleftrightarrow> (\<exists>S'. c = SKIP {S'} \<and> S \<le> S')" |
91 text_raw{*}%endsnip*} |
67 by (cases c) (auto simp:less_eq_acom_def anno_def) |
92 |
68 |
93 text_raw{*\snip{subbdef}{1}{1}{% *} |
69 lemma Assign_le[simp]: "x ::= e {S} \<le> c \<longleftrightarrow> (\<exists>S'. c = x ::= e {S'} \<and> S \<le> S')" |
94 fun sub\<^isub>2 :: "'a acom \<Rightarrow> 'a acom" where |
70 by (cases c) (auto simp:less_eq_acom_def anno_def) |
95 "sub\<^isub>2(C\<^isub>1;C\<^isub>2) = C\<^isub>2" | |
71 |
96 "sub\<^isub>2(IF b THEN {P\<^isub>1} C\<^isub>1 ELSE {P\<^isub>2} C\<^isub>2 {Q}) = C\<^isub>2" |
72 lemma Seq_le[simp]: "C1;C2 \<le> C \<longleftrightarrow> (\<exists>C1' C2'. C = C1';C2' \<and> C1 \<le> C1' \<and> C2 \<le> C2')" |
97 text_raw{*}%endsnip*} |
73 apply (cases C) |
98 |
74 apply(auto simp: less_eq_acom_annos list_all2_append size_annos_same2) |
99 text_raw{*\snip{annoadef}{1}{1}{% *} |
75 done |
100 fun anno\<^isub>1 :: "'a acom \<Rightarrow> 'a" where |
76 |
101 "anno\<^isub>1(IF b THEN {P\<^isub>1} C\<^isub>1 ELSE {P\<^isub>2} C\<^isub>2 {Q}) = P\<^isub>1" | |
77 lemma If_le[simp]: "IF b THEN {p1} C1 ELSE {p2} C2 {S} \<le> C \<longleftrightarrow> |
102 "anno\<^isub>1({I} WHILE b DO {P} C {Q}) = I" |
78 (\<exists>p1' p2' C1' C2' S'. C = IF b THEN {p1'} C1' ELSE {p2'} C2' {S'} \<and> |
103 text_raw{*}%endsnip*} |
79 p1 \<le> p1' \<and> p2 \<le> p2' \<and> C1 \<le> C1' \<and> C2 \<le> C2' \<and> S \<le> S')" |
104 |
80 apply (cases C) |
105 text_raw{*\snip{annobdef}{1}{2}{% *} |
81 apply(auto simp: less_eq_acom_annos list_all2_append size_annos_same2) |
106 fun anno\<^isub>2 :: "'a acom \<Rightarrow> 'a" where |
82 done |
107 "anno\<^isub>2(IF b THEN {P\<^isub>1} C\<^isub>1 ELSE {P\<^isub>2} C\<^isub>2 {Q}) = P\<^isub>2" | |
83 |
108 "anno\<^isub>2({I} WHILE b DO {P} C {Q}) = P" |
84 lemma While_le[simp]: "{I} WHILE b DO {p} C {P} \<le> W \<longleftrightarrow> |
109 text_raw{*}%endsnip*} |
85 (\<exists>I' p' C' P'. W = {I'} WHILE b DO {p'} C' {P'} \<and> C \<le> C' \<and> p \<le> p' \<and> I \<le> I' \<and> P \<le> P')" |
110 |
86 apply (cases W) |
111 fun merge :: "com \<Rightarrow> 'a acom set \<Rightarrow> 'a set acom" where |
87 apply(auto simp: less_eq_acom_annos list_all2_append size_annos_same2) |
112 "merge com.SKIP M = (SKIP {post ` M})" | |
88 done |
113 "merge (x ::= a) M = (x ::= a {post ` M})" | |
89 |
114 "merge (c1;c2) M = |
90 lemma mono_post: "C \<le> C' \<Longrightarrow> post C \<le> post C'" |
115 merge c1 (sub\<^isub>1 ` M); merge c2 (sub\<^isub>2 ` M)" | |
91 using annos_ne[of C'] |
116 "merge (IF b THEN c1 ELSE c2) M = |
92 by(auto simp: post_def less_eq_acom_def last_conv_nth[OF annos_ne] anno_def |
117 IF b THEN {anno\<^isub>1 ` M} merge c1 (sub\<^isub>1 ` M) ELSE {anno\<^isub>2 ` M} merge c2 (sub\<^isub>2 ` M) |
93 dest: size_annos_same) |
118 {post ` M}" | |
94 |
119 "merge (WHILE b DO c) M = |
95 definition Inf_acom :: "com \<Rightarrow> 'a::complete_lattice acom set \<Rightarrow> 'a acom" where |
120 {anno\<^isub>1 ` M} |
96 "Inf_acom c M = annotate (\<lambda>p. INF C:M. anno C p) c" |
121 WHILE b DO {anno\<^isub>2 ` M} merge c (sub\<^isub>1 ` M) |
|
122 {post ` M}" |
|
123 |
97 |
124 interpretation |
98 interpretation |
125 Complete_Lattice "{C. strip C = c}" "map_acom Inter o (merge c)" for c |
99 Complete_Lattice "{C. strip C = c}" "Inf_acom c" for c |
126 proof |
100 proof |
127 case goal1 |
101 case goal1 thus ?case |
128 have "a:A \<Longrightarrow> map_acom Inter (merge (strip a) A) \<le> a" |
102 by(auto simp: Inf_acom_def less_eq_acom_def size_annos intro:INF_lower) |
129 proof(induction a arbitrary: A) |
103 next |
130 case Seq from Seq.prems show ?case by(force intro!: Seq.IH) |
104 case goal2 thus ?case |
131 next |
105 by(auto simp: Inf_acom_def less_eq_acom_def size_annos intro:INF_greatest) |
132 case If from If.prems show ?case by(force intro!: If.IH) |
106 next |
133 next |
107 case goal3 thus ?case by(auto simp: Inf_acom_def) |
134 case While from While.prems show ?case by(force intro!: While.IH) |
108 qed |
135 qed force+ |
|
136 with goal1 show ?case by auto |
|
137 next |
|
138 case goal2 |
|
139 thus ?case |
|
140 proof(simp, induction b arbitrary: c A) |
|
141 case SKIP thus ?case by (force simp:SKIP_le) |
|
142 next |
|
143 case Assign thus ?case by (force simp:Assign_le) |
|
144 next |
|
145 case Seq from Seq.prems show ?case by(force intro!: Seq.IH simp:Seq_le) |
|
146 next |
|
147 case If from If.prems show ?case by (force simp: If_le intro!: If.IH) |
|
148 next |
|
149 case While from While.prems show ?case by(fastforce simp: While_le intro: While.IH) |
|
150 qed |
|
151 next |
|
152 case goal3 |
|
153 have "strip(merge c A) = c" |
|
154 proof(induction c arbitrary: A) |
|
155 case Seq from Seq.prems show ?case by (fastforce simp: strip_eq_Seq subset_iff intro!: Seq.IH) |
|
156 next |
|
157 case If from If.prems show ?case by (fastforce intro!: If.IH simp: strip_eq_If) |
|
158 next |
|
159 case While from While.prems show ?case by(fastforce intro: While.IH simp: strip_eq_While) |
|
160 qed auto |
|
161 thus ?case by auto |
|
162 qed |
|
163 |
|
164 lemma le_post: "c \<le> d \<Longrightarrow> post c \<le> post d" |
|
165 by(induction c d rule: less_eq_acom.induct) auto |
|
166 |
109 |
167 |
110 |
168 subsubsection "Collecting semantics" |
111 subsubsection "Collecting semantics" |
169 |
112 |
170 definition "step = Step (\<lambda>x e S. {s(x := aval e s) |s. s : S}) (\<lambda>b S. {s:S. bval b s})" |
113 definition "step = Step (\<lambda>x e S. {s(x := aval e s) |s. s : S}) (\<lambda>b S. {s:S. bval b s})" |
174 |
117 |
175 lemma mono2_Step: fixes C1 C2 :: "'a::semilattice_sup acom" |
118 lemma mono2_Step: fixes C1 C2 :: "'a::semilattice_sup acom" |
176 assumes "!!x e S1 S2. S1 \<le> S2 \<Longrightarrow> f x e S1 \<le> f x e S2" |
119 assumes "!!x e S1 S2. S1 \<le> S2 \<Longrightarrow> f x e S1 \<le> f x e S2" |
177 "!!b S1 S2. S1 \<le> S2 \<Longrightarrow> g b S1 \<le> g b S2" |
120 "!!b S1 S2. S1 \<le> S2 \<Longrightarrow> g b S1 \<le> g b S2" |
178 shows "C1 \<le> C2 \<Longrightarrow> S1 \<le> S2 \<Longrightarrow> Step f g S1 C1 \<le> Step f g S2 C2" |
121 shows "C1 \<le> C2 \<Longrightarrow> S1 \<le> S2 \<Longrightarrow> Step f g S1 C1 \<le> Step f g S2 C2" |
179 proof(induction C1 C2 arbitrary: S1 S2 rule: less_eq_acom.induct) |
122 proof(induction S1 C1 arbitrary: C2 S2 rule: Step.induct) |
180 case 2 thus ?case by (fastforce simp: assms(1)) |
123 case 1 thus ?case by(auto) |
181 next |
124 next |
182 case 3 thus ?case by(simp add: le_post) |
125 case 2 thus ?case by (auto simp: assms(1)) |
|
126 next |
|
127 case 3 thus ?case by(auto simp: mono_post) |
183 next |
128 next |
184 case 4 thus ?case |
129 case 4 thus ?case |
185 by(simp add: subset_iff assms(2)) (metis le_post le_supI1 le_supI2) |
130 by(auto simp: subset_iff assms(2)) |
|
131 (metis mono_post le_supI1 le_supI2)+ |
186 next |
132 next |
187 case 5 thus ?case |
133 case 5 thus ?case |
188 by(simp add: subset_iff assms(2)) (metis le_post le_supI1 le_supI2) |
134 by(auto simp: subset_iff assms(2)) |
189 qed auto |
135 (metis mono_post le_supI1 le_supI2)+ |
|
136 qed |
190 |
137 |
191 lemma mono2_step: "C1 \<le> C2 \<Longrightarrow> S1 \<subseteq> S2 \<Longrightarrow> step S1 C1 \<le> step S2 C2" |
138 lemma mono2_step: "C1 \<le> C2 \<Longrightarrow> S1 \<subseteq> S2 \<Longrightarrow> step S1 C1 \<le> step S2 C2" |
192 unfolding step_def by(rule mono2_Step) auto |
139 unfolding step_def by(rule mono2_Step) auto |
193 |
140 |
194 lemma mono_step: "mono (step S)" |
141 lemma mono_step: "mono (step S)" |
209 by(simp add: CS_def index_lfp[simplified]) |
156 by(simp add: CS_def index_lfp[simplified]) |
210 |
157 |
211 |
158 |
212 subsubsection "Relation to big-step semantics" |
159 subsubsection "Relation to big-step semantics" |
213 |
160 |
214 lemma post_merge: "\<forall> c' \<in> M. strip c' = c \<Longrightarrow> post (merge c M) = post ` M" |
161 lemma asize_nz: "asize(c::com) \<noteq> 0" |
215 proof(induction c arbitrary: M) |
162 by (metis length_0_conv length_annos_annotate annos_ne) |
216 case (Seq c1 c2) |
163 |
217 have "post ` M = post ` sub\<^isub>2 ` M" using Seq.prems by (force simp: strip_eq_Seq) |
164 lemma post_Inf_acom: |
218 moreover have "\<forall> c' \<in> sub\<^isub>2 ` M. strip c' = c2" using Seq.prems by (auto simp: strip_eq_Seq) |
165 "\<forall>C\<in>M. strip C = c \<Longrightarrow> post (Inf_acom c M) = \<Inter>(post ` M)" |
219 ultimately show ?case using Seq.IH(2)[of "sub\<^isub>2 ` M"] by simp |
166 apply(subgoal_tac "\<forall>C\<in>M. size(annos C) = asize c") |
220 qed simp_all |
167 apply(simp add: post_anno_asize Inf_acom_def asize_nz neq0_conv[symmetric]) |
221 |
168 apply(simp add: size_annos) |
|
169 done |
222 |
170 |
223 lemma post_lfp: "post(lfp c f) = (\<Inter>{post C|C. strip C = c \<and> f C \<le> C})" |
171 lemma post_lfp: "post(lfp c f) = (\<Inter>{post C|C. strip C = c \<and> f C \<le> C})" |
224 by(auto simp add: lfp_def post_merge) |
172 by(auto simp add: lfp_def post_Inf_acom) |
225 |
173 |
226 lemma big_step_post_step: |
174 lemma big_step_post_step: |
227 "\<lbrakk> (c, s) \<Rightarrow> t; strip C = c; s \<in> S; step S C \<le> C \<rbrakk> \<Longrightarrow> t \<in> post C" |
175 "\<lbrakk> (c, s) \<Rightarrow> t; strip C = c; s \<in> S; step S C \<le> C \<rbrakk> \<Longrightarrow> t \<in> post C" |
228 proof(induction arbitrary: C S rule: big_step_induct) |
176 proof(induction arbitrary: C S rule: big_step_induct) |
229 case Skip thus ?case by(auto simp: strip_eq_SKIP step_def) |
177 case Skip thus ?case by(auto simp: strip_eq_SKIP step_def post_def) |
230 next |
178 next |
231 case Assign thus ?case by(fastforce simp: strip_eq_Assign step_def) |
179 case Assign thus ?case |
232 next |
180 by(fastforce simp: strip_eq_Assign step_def post_def) |
233 case Seq thus ?case by(fastforce simp: strip_eq_Seq step_def) |
181 next |
234 next |
182 case Seq thus ?case |
235 case IfTrue thus ?case apply(auto simp: strip_eq_If step_def) |
183 by(fastforce simp: strip_eq_Seq step_def post_def last_append annos_ne) |
|
184 next |
|
185 case IfTrue thus ?case apply(auto simp: strip_eq_If step_def post_def) |
236 by (metis (lifting,full_types) mem_Collect_eq set_mp) |
186 by (metis (lifting,full_types) mem_Collect_eq set_mp) |
237 next |
187 next |
238 case IfFalse thus ?case apply(auto simp: strip_eq_If step_def) |
188 case IfFalse thus ?case apply(auto simp: strip_eq_If step_def post_def) |
239 by (metis (lifting,full_types) mem_Collect_eq set_mp) |
189 by (metis (lifting,full_types) mem_Collect_eq set_mp) |
240 next |
190 next |
241 case (WhileTrue b s1 c' s2 s3) |
191 case (WhileTrue b s1 c' s2 s3) |
242 from WhileTrue.prems(1) obtain I P C' Q where "C = {I} WHILE b DO {P} C' {Q}" "strip C' = c'" |
192 from WhileTrue.prems(1) obtain I P C' Q where "C = {I} WHILE b DO {P} C' {Q}" "strip C' = c'" |
243 by(auto simp: strip_eq_While) |
193 by(auto simp: strip_eq_While) |
244 from WhileTrue.prems(3) `C = _` |
194 from WhileTrue.prems(3) `C = _` |
245 have "step P C' \<le> C'" "{s \<in> I. bval b s} \<le> P" "S \<le> I" "step (post C') C \<le> C" |
195 have "step P C' \<le> C'" "{s \<in> I. bval b s} \<le> P" "S \<le> I" "step (post C') C \<le> C" |
246 by (auto simp: step_def) |
196 by (auto simp: step_def post_def) |
247 have "step {s \<in> I. bval b s} C' \<le> C'" |
197 have "step {s \<in> I. bval b s} C' \<le> C'" |
248 by (rule order_trans[OF mono2_step[OF order_refl `{s \<in> I. bval b s} \<le> P`] `step P C' \<le> C'`]) |
198 by (rule order_trans[OF mono2_step[OF order_refl `{s \<in> I. bval b s} \<le> P`] `step P C' \<le> C'`]) |
249 have "s1: {s:I. bval b s}" using `s1 \<in> S` `S \<subseteq> I` `bval b s1` by auto |
199 have "s1: {s:I. bval b s}" using `s1 \<in> S` `S \<subseteq> I` `bval b s1` by auto |
250 note s2_in_post_C' = WhileTrue.IH(1)[OF `strip C' = c'` this `step {s \<in> I. bval b s} C' \<le> C'`] |
200 note s2_in_post_C' = WhileTrue.IH(1)[OF `strip C' = c'` this `step {s \<in> I. bval b s} C' \<le> C'`] |
251 from WhileTrue.IH(2)[OF WhileTrue.prems(1) s2_in_post_C' `step (post C') C \<le> C`] |
201 from WhileTrue.IH(2)[OF WhileTrue.prems(1) s2_in_post_C' `step (post C') C \<le> C`] |
252 show ?case . |
202 show ?case . |
253 next |
203 next |
254 case (WhileFalse b s1 c') thus ?case by (force simp: strip_eq_While step_def) |
204 case (WhileFalse b s1 c') thus ?case |
|
205 by (force simp: strip_eq_While step_def post_def) |
255 qed |
206 qed |
256 |
207 |
257 lemma big_step_lfp: "\<lbrakk> (c,s) \<Rightarrow> t; s \<in> S \<rbrakk> \<Longrightarrow> t \<in> post(lfp c (step S))" |
208 lemma big_step_lfp: "\<lbrakk> (c,s) \<Rightarrow> t; s \<in> S \<rbrakk> \<Longrightarrow> t \<in> post(lfp c (step S))" |
258 by(auto simp add: post_lfp intro: big_step_post_step) |
209 by(auto simp add: post_lfp intro: big_step_post_step) |
259 |
210 |