author | wenzelm |
Fri, 03 May 2019 20:03:45 +0200 (2019-05-03) | |
changeset 70244 | 2ca87b481077 |
parent 69712 | dc85b5b3a532 |
child 80914 | d97fdabd9e2b |
permissions | -rw-r--r-- |
68778 | 1 |
(* Author: Tobias Nipkow *) |
2 |
||
3 |
subsection "Collecting Semantics of Commands" |
|
4 |
||
45623
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
diff
changeset
|
5 |
theory Collecting |
49487 | 6 |
imports Complete_Lattice Big_Step ACom |
45623
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
diff
changeset
|
7 |
begin |
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
diff
changeset
|
8 |
|
68778 | 9 |
subsubsection "The generic Step function" |
51389 | 10 |
|
11 |
notation |
|
12 |
sup (infixl "\<squnion>" 65) and |
|
13 |
inf (infixl "\<sqinter>" 70) and |
|
14 |
bot ("\<bottom>") and |
|
15 |
top ("\<top>") |
|
16 |
||
52019
a4cbca8f7342
finally: acom with pointwise access and update of annotations
nipkow
parents:
51630
diff
changeset
|
17 |
context |
a4cbca8f7342
finally: acom with pointwise access and update of annotations
nipkow
parents:
51630
diff
changeset
|
18 |
fixes f :: "vname \<Rightarrow> aexp \<Rightarrow> 'a \<Rightarrow> 'a::sup" |
a4cbca8f7342
finally: acom with pointwise access and update of annotations
nipkow
parents:
51630
diff
changeset
|
19 |
fixes g :: "bexp \<Rightarrow> 'a \<Rightarrow> 'a" |
a4cbca8f7342
finally: acom with pointwise access and update of annotations
nipkow
parents:
51630
diff
changeset
|
20 |
begin |
a4cbca8f7342
finally: acom with pointwise access and update of annotations
nipkow
parents:
51630
diff
changeset
|
21 |
fun Step :: "'a \<Rightarrow> 'a acom \<Rightarrow> 'a acom" where |
a4cbca8f7342
finally: acom with pointwise access and update of annotations
nipkow
parents:
51630
diff
changeset
|
22 |
"Step S (SKIP {Q}) = (SKIP {S})" | |
a4cbca8f7342
finally: acom with pointwise access and update of annotations
nipkow
parents:
51630
diff
changeset
|
23 |
"Step S (x ::= e {Q}) = |
51390 | 24 |
x ::= e {f x e S}" | |
52046
bc01725d7918
replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
nipkow
parents:
52019
diff
changeset
|
25 |
"Step S (C1;; C2) = Step S C1;; Step (post C1) C2" | |
52019
a4cbca8f7342
finally: acom with pointwise access and update of annotations
nipkow
parents:
51630
diff
changeset
|
26 |
"Step S (IF b THEN {P1} C1 ELSE {P2} C2 {Q}) = |
a4cbca8f7342
finally: acom with pointwise access and update of annotations
nipkow
parents:
51630
diff
changeset
|
27 |
IF b THEN {g b S} Step P1 C1 ELSE {g (Not b) S} Step P2 C2 |
51390 | 28 |
{post C1 \<squnion> post C2}" | |
52019
a4cbca8f7342
finally: acom with pointwise access and update of annotations
nipkow
parents:
51630
diff
changeset
|
29 |
"Step S ({I} WHILE b DO {P} C {Q}) = |
a4cbca8f7342
finally: acom with pointwise access and update of annotations
nipkow
parents:
51630
diff
changeset
|
30 |
{S \<squnion> post C} WHILE b DO {g b I} Step P C {g (Not b) I}" |
a4cbca8f7342
finally: acom with pointwise access and update of annotations
nipkow
parents:
51630
diff
changeset
|
31 |
end |
51389 | 32 |
|
51390 | 33 |
lemma strip_Step[simp]: "strip(Step f g S C) = strip C" |
34 |
by(induct C arbitrary: S) auto |
|
51389 | 35 |
|
36 |
||
45623
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
diff
changeset
|
37 |
subsubsection "Annotated commands as a complete lattice" |
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
diff
changeset
|
38 |
|
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
diff
changeset
|
39 |
instantiation acom :: (order) order |
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
diff
changeset
|
40 |
begin |
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
diff
changeset
|
41 |
|
52019
a4cbca8f7342
finally: acom with pointwise access and update of annotations
nipkow
parents:
51630
diff
changeset
|
42 |
definition less_eq_acom :: "('a::order)acom \<Rightarrow> 'a acom \<Rightarrow> bool" where |
a4cbca8f7342
finally: acom with pointwise access and update of annotations
nipkow
parents:
51630
diff
changeset
|
43 |
"C1 \<le> C2 \<longleftrightarrow> strip C1 = strip C2 \<and> (\<forall>p<size(annos C1). anno C1 p \<le> anno C2 p)" |
45623
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
diff
changeset
|
44 |
|
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
diff
changeset
|
45 |
definition less_acom :: "'a acom \<Rightarrow> 'a acom \<Rightarrow> bool" where |
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
diff
changeset
|
46 |
"less_acom x y = (x \<le> y \<and> \<not> y \<le> x)" |
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
diff
changeset
|
47 |
|
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
diff
changeset
|
48 |
instance |
61179 | 49 |
proof (standard, goal_cases) |
50 |
case 1 show ?case by(simp add: less_acom_def) |
|
45623
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
diff
changeset
|
51 |
next |
61179 | 52 |
case 2 thus ?case by(auto simp: less_eq_acom_def) |
45623
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
diff
changeset
|
53 |
next |
61179 | 54 |
case 3 thus ?case by(fastforce simp: less_eq_acom_def size_annos) |
45623
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
diff
changeset
|
55 |
next |
61179 | 56 |
case 4 thus ?case |
52019
a4cbca8f7342
finally: acom with pointwise access and update of annotations
nipkow
parents:
51630
diff
changeset
|
57 |
by(fastforce simp: le_antisym less_eq_acom_def size_annos |
a4cbca8f7342
finally: acom with pointwise access and update of annotations
nipkow
parents:
51630
diff
changeset
|
58 |
eq_acom_iff_strip_anno) |
45623
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
diff
changeset
|
59 |
qed |
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
diff
changeset
|
60 |
|
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
diff
changeset
|
61 |
end |
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
diff
changeset
|
62 |
|
52019
a4cbca8f7342
finally: acom with pointwise access and update of annotations
nipkow
parents:
51630
diff
changeset
|
63 |
lemma less_eq_acom_annos: |
67399 | 64 |
"C1 \<le> C2 \<longleftrightarrow> strip C1 = strip C2 \<and> list_all2 (\<le>) (annos C1) (annos C2)" |
52019
a4cbca8f7342
finally: acom with pointwise access and update of annotations
nipkow
parents:
51630
diff
changeset
|
65 |
by(auto simp add: less_eq_acom_def anno_def list_all2_conv_all_nth size_annos_same2) |
a4cbca8f7342
finally: acom with pointwise access and update of annotations
nipkow
parents:
51630
diff
changeset
|
66 |
|
a4cbca8f7342
finally: acom with pointwise access and update of annotations
nipkow
parents:
51630
diff
changeset
|
67 |
lemma SKIP_le[simp]: "SKIP {S} \<le> c \<longleftrightarrow> (\<exists>S'. c = SKIP {S'} \<and> S \<le> S')" |
a4cbca8f7342
finally: acom with pointwise access and update of annotations
nipkow
parents:
51630
diff
changeset
|
68 |
by (cases c) (auto simp:less_eq_acom_def anno_def) |
45903 | 69 |
|
52019
a4cbca8f7342
finally: acom with pointwise access and update of annotations
nipkow
parents:
51630
diff
changeset
|
70 |
lemma Assign_le[simp]: "x ::= e {S} \<le> c \<longleftrightarrow> (\<exists>S'. c = x ::= e {S'} \<and> S \<le> S')" |
a4cbca8f7342
finally: acom with pointwise access and update of annotations
nipkow
parents:
51630
diff
changeset
|
71 |
by (cases c) (auto simp:less_eq_acom_def anno_def) |
45903 | 72 |
|
52046
bc01725d7918
replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
nipkow
parents:
52019
diff
changeset
|
73 |
lemma Seq_le[simp]: "C1;;C2 \<le> C \<longleftrightarrow> (\<exists>C1' C2'. C = C1';;C2' \<and> C1 \<le> C1' \<and> C2 \<le> C2')" |
52019
a4cbca8f7342
finally: acom with pointwise access and update of annotations
nipkow
parents:
51630
diff
changeset
|
74 |
apply (cases C) |
a4cbca8f7342
finally: acom with pointwise access and update of annotations
nipkow
parents:
51630
diff
changeset
|
75 |
apply(auto simp: less_eq_acom_annos list_all2_append size_annos_same2) |
a4cbca8f7342
finally: acom with pointwise access and update of annotations
nipkow
parents:
51630
diff
changeset
|
76 |
done |
49095 | 77 |
|
52019
a4cbca8f7342
finally: acom with pointwise access and update of annotations
nipkow
parents:
51630
diff
changeset
|
78 |
lemma If_le[simp]: "IF b THEN {p1} C1 ELSE {p2} C2 {S} \<le> C \<longleftrightarrow> |
a4cbca8f7342
finally: acom with pointwise access and update of annotations
nipkow
parents:
51630
diff
changeset
|
79 |
(\<exists>p1' p2' C1' C2' S'. C = IF b THEN {p1'} C1' ELSE {p2'} C2' {S'} \<and> |
a4cbca8f7342
finally: acom with pointwise access and update of annotations
nipkow
parents:
51630
diff
changeset
|
80 |
p1 \<le> p1' \<and> p2 \<le> p2' \<and> C1 \<le> C1' \<and> C2 \<le> C2' \<and> S \<le> S')" |
a4cbca8f7342
finally: acom with pointwise access and update of annotations
nipkow
parents:
51630
diff
changeset
|
81 |
apply (cases C) |
a4cbca8f7342
finally: acom with pointwise access and update of annotations
nipkow
parents:
51630
diff
changeset
|
82 |
apply(auto simp: less_eq_acom_annos list_all2_append size_annos_same2) |
a4cbca8f7342
finally: acom with pointwise access and update of annotations
nipkow
parents:
51630
diff
changeset
|
83 |
done |
45903 | 84 |
|
52019
a4cbca8f7342
finally: acom with pointwise access and update of annotations
nipkow
parents:
51630
diff
changeset
|
85 |
lemma While_le[simp]: "{I} WHILE b DO {p} C {P} \<le> W \<longleftrightarrow> |
a4cbca8f7342
finally: acom with pointwise access and update of annotations
nipkow
parents:
51630
diff
changeset
|
86 |
(\<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')" |
a4cbca8f7342
finally: acom with pointwise access and update of annotations
nipkow
parents:
51630
diff
changeset
|
87 |
apply (cases W) |
a4cbca8f7342
finally: acom with pointwise access and update of annotations
nipkow
parents:
51630
diff
changeset
|
88 |
apply(auto simp: less_eq_acom_annos list_all2_append size_annos_same2) |
a4cbca8f7342
finally: acom with pointwise access and update of annotations
nipkow
parents:
51630
diff
changeset
|
89 |
done |
a4cbca8f7342
finally: acom with pointwise access and update of annotations
nipkow
parents:
51630
diff
changeset
|
90 |
|
a4cbca8f7342
finally: acom with pointwise access and update of annotations
nipkow
parents:
51630
diff
changeset
|
91 |
lemma mono_post: "C \<le> C' \<Longrightarrow> post C \<le> post C'" |
a4cbca8f7342
finally: acom with pointwise access and update of annotations
nipkow
parents:
51630
diff
changeset
|
92 |
using annos_ne[of C'] |
a4cbca8f7342
finally: acom with pointwise access and update of annotations
nipkow
parents:
51630
diff
changeset
|
93 |
by(auto simp: post_def less_eq_acom_def last_conv_nth[OF annos_ne] anno_def |
a4cbca8f7342
finally: acom with pointwise access and update of annotations
nipkow
parents:
51630
diff
changeset
|
94 |
dest: size_annos_same) |
a4cbca8f7342
finally: acom with pointwise access and update of annotations
nipkow
parents:
51630
diff
changeset
|
95 |
|
a4cbca8f7342
finally: acom with pointwise access and update of annotations
nipkow
parents:
51630
diff
changeset
|
96 |
definition Inf_acom :: "com \<Rightarrow> 'a::complete_lattice acom set \<Rightarrow> 'a acom" where |
69260
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
haftmann
parents:
68778
diff
changeset
|
97 |
"Inf_acom c M = annotate (\<lambda>p. INF C\<in>M. anno C p) c" |
45623
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
diff
changeset
|
98 |
|
61890
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
haftmann
parents:
61670
diff
changeset
|
99 |
global_interpretation |
52019
a4cbca8f7342
finally: acom with pointwise access and update of annotations
nipkow
parents:
51630
diff
changeset
|
100 |
Complete_Lattice "{C. strip C = c}" "Inf_acom c" for c |
61179 | 101 |
proof (standard, goal_cases) |
102 |
case 1 thus ?case |
|
52019
a4cbca8f7342
finally: acom with pointwise access and update of annotations
nipkow
parents:
51630
diff
changeset
|
103 |
by(auto simp: Inf_acom_def less_eq_acom_def size_annos intro:INF_lower) |
45623
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
diff
changeset
|
104 |
next |
61179 | 105 |
case 2 thus ?case |
52019
a4cbca8f7342
finally: acom with pointwise access and update of annotations
nipkow
parents:
51630
diff
changeset
|
106 |
by(auto simp: Inf_acom_def less_eq_acom_def size_annos intro:INF_greatest) |
45623
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
diff
changeset
|
107 |
next |
61179 | 108 |
case 3 thus ?case by(auto simp: Inf_acom_def) |
45623
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
diff
changeset
|
109 |
qed |
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
diff
changeset
|
110 |
|
49487 | 111 |
|
45623
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
diff
changeset
|
112 |
subsubsection "Collecting semantics" |
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
diff
changeset
|
113 |
|
67613 | 114 |
definition "step = Step (\<lambda>x e S. {s(x := aval e s) |s. s \<in> S}) (\<lambda>b S. {s:S. bval b s})" |
45623
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
diff
changeset
|
115 |
|
46070 | 116 |
definition CS :: "com \<Rightarrow> state set acom" where |
48759
ff570720ba1c
Improved complete lattice formalisation - no more index set.
nipkow
parents:
47818
diff
changeset
|
117 |
"CS c = lfp c (step UNIV)" |
45623
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
diff
changeset
|
118 |
|
51390 | 119 |
lemma mono2_Step: fixes C1 C2 :: "'a::semilattice_sup acom" |
120 |
assumes "!!x e S1 S2. S1 \<le> S2 \<Longrightarrow> f x e S1 \<le> f x e S2" |
|
121 |
"!!b S1 S2. S1 \<le> S2 \<Longrightarrow> g b S1 \<le> g b S2" |
|
122 |
shows "C1 \<le> C2 \<Longrightarrow> S1 \<le> S2 \<Longrightarrow> Step f g S1 C1 \<le> Step f g S2 C2" |
|
52019
a4cbca8f7342
finally: acom with pointwise access and update of annotations
nipkow
parents:
51630
diff
changeset
|
123 |
proof(induction S1 C1 arbitrary: C2 S2 rule: Step.induct) |
a4cbca8f7342
finally: acom with pointwise access and update of annotations
nipkow
parents:
51630
diff
changeset
|
124 |
case 1 thus ?case by(auto) |
45623
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
diff
changeset
|
125 |
next |
52019
a4cbca8f7342
finally: acom with pointwise access and update of annotations
nipkow
parents:
51630
diff
changeset
|
126 |
case 2 thus ?case by (auto simp: assms(1)) |
a4cbca8f7342
finally: acom with pointwise access and update of annotations
nipkow
parents:
51630
diff
changeset
|
127 |
next |
a4cbca8f7342
finally: acom with pointwise access and update of annotations
nipkow
parents:
51630
diff
changeset
|
128 |
case 3 thus ?case by(auto simp: mono_post) |
45623
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
diff
changeset
|
129 |
next |
51390 | 130 |
case 4 thus ?case |
52019
a4cbca8f7342
finally: acom with pointwise access and update of annotations
nipkow
parents:
51630
diff
changeset
|
131 |
by(auto simp: subset_iff assms(2)) |
a4cbca8f7342
finally: acom with pointwise access and update of annotations
nipkow
parents:
51630
diff
changeset
|
132 |
(metis mono_post le_supI1 le_supI2)+ |
45623
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
diff
changeset
|
133 |
next |
51390 | 134 |
case 5 thus ?case |
52019
a4cbca8f7342
finally: acom with pointwise access and update of annotations
nipkow
parents:
51630
diff
changeset
|
135 |
by(auto simp: subset_iff assms(2)) |
a4cbca8f7342
finally: acom with pointwise access and update of annotations
nipkow
parents:
51630
diff
changeset
|
136 |
(metis mono_post le_supI1 le_supI2)+ |
a4cbca8f7342
finally: acom with pointwise access and update of annotations
nipkow
parents:
51630
diff
changeset
|
137 |
qed |
45623
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
diff
changeset
|
138 |
|
51390 | 139 |
lemma mono2_step: "C1 \<le> C2 \<Longrightarrow> S1 \<subseteq> S2 \<Longrightarrow> step S1 C1 \<le> step S2 C2" |
140 |
unfolding step_def by(rule mono2_Step) auto |
|
141 |
||
45655
a49f9428aba4
simplified Collecting1 and renamed: step -> step', step_cs -> step
nipkow
parents:
45623
diff
changeset
|
142 |
lemma mono_step: "mono (step S)" |
46334 | 143 |
by(blast intro: monoI mono2_step) |
45623
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
diff
changeset
|
144 |
|
49095 | 145 |
lemma strip_step: "strip(step S C) = strip C" |
51390 | 146 |
by (induction C arbitrary: S) (auto simp: step_def) |
45623
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
diff
changeset
|
147 |
|
48759
ff570720ba1c
Improved complete lattice formalisation - no more index set.
nipkow
parents:
47818
diff
changeset
|
148 |
lemma lfp_cs_unfold: "lfp c (step S) = step S (lfp c (step S))" |
45903 | 149 |
apply(rule lfp_unfold[OF _ mono_step]) |
150 |
apply(simp add: strip_step) |
|
151 |
done |
|
45623
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
diff
changeset
|
152 |
|
46070 | 153 |
lemma CS_unfold: "CS c = step UNIV (CS c)" |
45623
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
diff
changeset
|
154 |
by (metis CS_def lfp_cs_unfold) |
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
diff
changeset
|
155 |
|
46070 | 156 |
lemma strip_CS[simp]: "strip(CS c) = c" |
45903 | 157 |
by(simp add: CS_def index_lfp[simplified]) |
45623
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
diff
changeset
|
158 |
|
49487 | 159 |
|
160 |
subsubsection "Relation to big-step semantics" |
|
161 |
||
52019
a4cbca8f7342
finally: acom with pointwise access and update of annotations
nipkow
parents:
51630
diff
changeset
|
162 |
lemma asize_nz: "asize(c::com) \<noteq> 0" |
a4cbca8f7342
finally: acom with pointwise access and update of annotations
nipkow
parents:
51630
diff
changeset
|
163 |
by (metis length_0_conv length_annos_annotate annos_ne) |
49487 | 164 |
|
52019
a4cbca8f7342
finally: acom with pointwise access and update of annotations
nipkow
parents:
51630
diff
changeset
|
165 |
lemma post_Inf_acom: |
a4cbca8f7342
finally: acom with pointwise access and update of annotations
nipkow
parents:
51630
diff
changeset
|
166 |
"\<forall>C\<in>M. strip C = c \<Longrightarrow> post (Inf_acom c M) = \<Inter>(post ` M)" |
a4cbca8f7342
finally: acom with pointwise access and update of annotations
nipkow
parents:
51630
diff
changeset
|
167 |
apply(subgoal_tac "\<forall>C\<in>M. size(annos C) = asize c") |
a4cbca8f7342
finally: acom with pointwise access and update of annotations
nipkow
parents:
51630
diff
changeset
|
168 |
apply(simp add: post_anno_asize Inf_acom_def asize_nz neq0_conv[symmetric]) |
a4cbca8f7342
finally: acom with pointwise access and update of annotations
nipkow
parents:
51630
diff
changeset
|
169 |
apply(simp add: size_annos) |
a4cbca8f7342
finally: acom with pointwise access and update of annotations
nipkow
parents:
51630
diff
changeset
|
170 |
done |
49487 | 171 |
|
172 |
lemma post_lfp: "post(lfp c f) = (\<Inter>{post C|C. strip C = c \<and> f C \<le> C})" |
|
52019
a4cbca8f7342
finally: acom with pointwise access and update of annotations
nipkow
parents:
51630
diff
changeset
|
173 |
by(auto simp add: lfp_def post_Inf_acom) |
49487 | 174 |
|
175 |
lemma big_step_post_step: |
|
176 |
"\<lbrakk> (c, s) \<Rightarrow> t; strip C = c; s \<in> S; step S C \<le> C \<rbrakk> \<Longrightarrow> t \<in> post C" |
|
177 |
proof(induction arbitrary: C S rule: big_step_induct) |
|
52019
a4cbca8f7342
finally: acom with pointwise access and update of annotations
nipkow
parents:
51630
diff
changeset
|
178 |
case Skip thus ?case by(auto simp: strip_eq_SKIP step_def post_def) |
49487 | 179 |
next |
52019
a4cbca8f7342
finally: acom with pointwise access and update of annotations
nipkow
parents:
51630
diff
changeset
|
180 |
case Assign thus ?case |
a4cbca8f7342
finally: acom with pointwise access and update of annotations
nipkow
parents:
51630
diff
changeset
|
181 |
by(fastforce simp: strip_eq_Assign step_def post_def) |
49487 | 182 |
next |
52019
a4cbca8f7342
finally: acom with pointwise access and update of annotations
nipkow
parents:
51630
diff
changeset
|
183 |
case Seq thus ?case |
a4cbca8f7342
finally: acom with pointwise access and update of annotations
nipkow
parents:
51630
diff
changeset
|
184 |
by(fastforce simp: strip_eq_Seq step_def post_def last_append annos_ne) |
49487 | 185 |
next |
52019
a4cbca8f7342
finally: acom with pointwise access and update of annotations
nipkow
parents:
51630
diff
changeset
|
186 |
case IfTrue thus ?case apply(auto simp: strip_eq_If step_def post_def) |
69712 | 187 |
by (metis (lifting,full_types) mem_Collect_eq subsetD) |
49487 | 188 |
next |
52019
a4cbca8f7342
finally: acom with pointwise access and update of annotations
nipkow
parents:
51630
diff
changeset
|
189 |
case IfFalse thus ?case apply(auto simp: strip_eq_If step_def post_def) |
69712 | 190 |
by (metis (lifting,full_types) mem_Collect_eq subsetD) |
49487 | 191 |
next |
192 |
case (WhileTrue b s1 c' s2 s3) |
|
193 |
from WhileTrue.prems(1) obtain I P C' Q where "C = {I} WHILE b DO {P} C' {Q}" "strip C' = c'" |
|
194 |
by(auto simp: strip_eq_While) |
|
67406 | 195 |
from WhileTrue.prems(3) \<open>C = _\<close> |
51390 | 196 |
have "step P C' \<le> C'" "{s \<in> I. bval b s} \<le> P" "S \<le> I" "step (post C') C \<le> C" |
52019
a4cbca8f7342
finally: acom with pointwise access and update of annotations
nipkow
parents:
51630
diff
changeset
|
197 |
by (auto simp: step_def post_def) |
49487 | 198 |
have "step {s \<in> I. bval b s} C' \<le> C'" |
67406 | 199 |
by (rule order_trans[OF mono2_step[OF order_refl \<open>{s \<in> I. bval b s} \<le> P\<close>] \<open>step P C' \<le> C'\<close>]) |
67613 | 200 |
have "s1 \<in> {s\<in>I. bval b s}" using \<open>s1 \<in> S\<close> \<open>S \<subseteq> I\<close> \<open>bval b s1\<close> by auto |
67406 | 201 |
note s2_in_post_C' = WhileTrue.IH(1)[OF \<open>strip C' = c'\<close> this \<open>step {s \<in> I. bval b s} C' \<le> C'\<close>] |
202 |
from WhileTrue.IH(2)[OF WhileTrue.prems(1) s2_in_post_C' \<open>step (post C') C \<le> C\<close>] |
|
49487 | 203 |
show ?case . |
204 |
next |
|
52019
a4cbca8f7342
finally: acom with pointwise access and update of annotations
nipkow
parents:
51630
diff
changeset
|
205 |
case (WhileFalse b s1 c') thus ?case |
a4cbca8f7342
finally: acom with pointwise access and update of annotations
nipkow
parents:
51630
diff
changeset
|
206 |
by (force simp: strip_eq_While step_def post_def) |
49487 | 207 |
qed |
208 |
||
209 |
lemma big_step_lfp: "\<lbrakk> (c,s) \<Rightarrow> t; s \<in> S \<rbrakk> \<Longrightarrow> t \<in> post(lfp c (step S))" |
|
210 |
by(auto simp add: post_lfp intro: big_step_post_step) |
|
211 |
||
67613 | 212 |
lemma big_step_CS: "(c,s) \<Rightarrow> t \<Longrightarrow> t \<in> post(CS c)" |
49487 | 213 |
by(simp add: CS_def big_step_lfp) |
214 |
||
45623
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
diff
changeset
|
215 |
end |