| author | nipkow | 
| Sun, 10 Mar 2013 18:29:10 +0100 | |
| changeset 51389 | 8a9f0503b1c0 | 
| parent 51039 | 3775bf0ea5b8 | 
| child 51390 | 1dff81cf425b | 
| permissions | -rw-r--r-- | 
| 
45623
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents:  
diff
changeset
 | 
1  | 
theory Collecting  | 
| 49487 | 2  | 
imports Complete_Lattice Big_Step ACom  | 
| 51389 | 3  | 
"~~/src/HOL/ex/Interpretation_with_Defs"  | 
| 
45623
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents:  
diff
changeset
 | 
4  | 
begin  | 
| 
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents:  
diff
changeset
 | 
5  | 
|
| 51389 | 6  | 
subsection "The generic Step function"  | 
7  | 
||
8  | 
notation  | 
|
9  | 
sup (infixl "\<squnion>" 65) and  | 
|
10  | 
inf (infixl "\<sqinter>" 70) and  | 
|
11  | 
  bot ("\<bottom>") and
 | 
|
12  | 
  top ("\<top>")
 | 
|
13  | 
||
14  | 
locale Step =  | 
|
15  | 
fixes step_assign :: "vname \<Rightarrow> aexp \<Rightarrow> 'a \<Rightarrow> 'a::sup"  | 
|
16  | 
and step_cond :: "bexp \<Rightarrow> 'a \<Rightarrow> 'a"  | 
|
17  | 
begin  | 
|
18  | 
||
19  | 
fun Step :: "'a \<Rightarrow> 'a acom \<Rightarrow> 'a acom" where  | 
|
20  | 
"Step S (SKIP {Q}) = (SKIP {S})" |
 | 
|
21  | 
"Step S (x ::= e {Q}) =
 | 
|
22  | 
  x ::= e {step_assign x e S}" |
 | 
|
23  | 
"Step S (C1; C2) = Step S C1; Step (post C1) C2" |  | 
|
24  | 
"Step S (IF b THEN {P1} C1 ELSE {P2} C2 {Q}) =
 | 
|
25  | 
  IF b THEN {step_cond b S} Step P1 C1 ELSE {step_cond (Not b) S} Step P2 C2
 | 
|
26  | 
  {post C1 \<squnion> post C2}" |
 | 
|
27  | 
"Step S ({I} WHILE b DO {P} C {Q}) =
 | 
|
28  | 
  {S \<squnion> post C} WHILE b DO {step_cond b I} Step P C {step_cond (Not b) I}"
 | 
|
29  | 
||
30  | 
end  | 
|
31  | 
||
32  | 
||
| 
45623
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents:  
diff
changeset
 | 
33  | 
subsection "Collecting Semantics of Commands"  | 
| 
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents:  
diff
changeset
 | 
34  | 
|
| 
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents:  
diff
changeset
 | 
35  | 
subsubsection "Annotated commands as a complete lattice"  | 
| 
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents:  
diff
changeset
 | 
36  | 
|
| 
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents:  
diff
changeset
 | 
37  | 
(* Orderings could also be lifted generically (thus subsuming the  | 
| 
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents:  
diff
changeset
 | 
38  | 
instantiation for preord and order), but then less_eq_acom would need to  | 
| 
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents:  
diff
changeset
 | 
39  | 
become a definition, eg less_eq_acom = lift2 less_eq, and then proofs would  | 
| 
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents:  
diff
changeset
 | 
40  | 
need to unfold this defn first. *)  | 
| 
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents:  
diff
changeset
 | 
41  | 
|
| 
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents:  
diff
changeset
 | 
42  | 
instantiation acom :: (order) order  | 
| 
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents:  
diff
changeset
 | 
43  | 
begin  | 
| 
 
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  | 
fun less_eq_acom :: "('a::order)acom \<Rightarrow> 'a acom \<Rightarrow> bool" where
 | 
| 49344 | 46  | 
"(SKIP {P}) \<le> (SKIP {P'}) = (P \<le> P')" |
 | 
47  | 
"(x ::= e {P}) \<le> (x' ::= e' {P'}) = (x=x' \<and> e=e' \<and> P \<le> P')" |
 | 
|
48  | 
"(C1;C2) \<le> (C1';C2') = (C1 \<le> C1' \<and> C2 \<le> C2')" |  | 
|
49  | 
"(IF b THEN {P1} C1 ELSE {P2} C2 {Q}) \<le> (IF b' THEN {P1'} C1' ELSE {P2'} C2' {Q'}) =
 | 
|
50  | 
(b=b' \<and> P1 \<le> P1' \<and> C1 \<le> C1' \<and> P2 \<le> P2' \<and> C2 \<le> C2' \<and> Q \<le> Q')" |  | 
|
51  | 
"({I} WHILE b DO {P} C {Q}) \<le> ({I'} WHILE b' DO {P'} C' {Q'}) =
 | 
|
52  | 
(b=b' \<and> C \<le> C' \<and> I \<le> I' \<and> P \<le> P' \<and> Q \<le> Q')" |  | 
|
| 
45623
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents:  
diff
changeset
 | 
53  | 
"less_eq_acom _ _ = False"  | 
| 
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents:  
diff
changeset
 | 
54  | 
|
| 
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents:  
diff
changeset
 | 
55  | 
lemma SKIP_le: "SKIP {S} \<le> c \<longleftrightarrow> (\<exists>S'. c = SKIP {S'} \<and> S \<le> S')"
 | 
| 
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents:  
diff
changeset
 | 
56  | 
by (cases c) auto  | 
| 
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents:  
diff
changeset
 | 
57  | 
|
| 
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents:  
diff
changeset
 | 
58  | 
lemma Assign_le: "x ::= e {S} \<le> c \<longleftrightarrow> (\<exists>S'. c = x ::= e {S'} \<and> S \<le> S')"
 | 
| 
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents:  
diff
changeset
 | 
59  | 
by (cases c) auto  | 
| 
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents:  
diff
changeset
 | 
60  | 
|
| 49095 | 61  | 
lemma Seq_le: "C1;C2 \<le> C \<longleftrightarrow> (\<exists>C1' C2'. C = C1';C2' \<and> C1 \<le> C1' \<and> C2 \<le> C2')"  | 
62  | 
by (cases C) auto  | 
|
| 
45623
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents:  
diff
changeset
 | 
63  | 
|
| 49095 | 64  | 
lemma If_le: "IF b THEN {p1} C1 ELSE {p2} C2 {S} \<le> C \<longleftrightarrow>
 | 
65  | 
  (\<exists>p1' p2' C1' C2' S'. C = IF b THEN {p1'} C1' ELSE {p2'} C2' {S'} \<and>
 | 
|
66  | 
p1 \<le> p1' \<and> p2 \<le> p2' \<and> C1 \<le> C1' \<and> C2 \<le> C2' \<and> S \<le> S')"  | 
|
67  | 
by (cases C) auto  | 
|
| 
45623
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents:  
diff
changeset
 | 
68  | 
|
| 49095 | 69  | 
lemma While_le: "{I} WHILE b DO {p} C {P} \<le> W \<longleftrightarrow>
 | 
70  | 
  (\<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')"
 | 
|
71  | 
by (cases W) auto  | 
|
| 
45623
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents:  
diff
changeset
 | 
72  | 
|
| 
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents:  
diff
changeset
 | 
73  | 
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
 | 
74  | 
"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
 | 
75  | 
|
| 
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents:  
diff
changeset
 | 
76  | 
instance  | 
| 
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents:  
diff
changeset
 | 
77  | 
proof  | 
| 
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents:  
diff
changeset
 | 
78  | 
case goal1 show ?case by(simp add: less_acom_def)  | 
| 
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents:  
diff
changeset
 | 
79  | 
next  | 
| 
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents:  
diff
changeset
 | 
80  | 
case goal2 thus ?case by (induct x) auto  | 
| 
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents:  
diff
changeset
 | 
81  | 
next  | 
| 
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents:  
diff
changeset
 | 
82  | 
case goal3 thus ?case  | 
| 
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents:  
diff
changeset
 | 
83  | 
apply(induct x y arbitrary: z rule: less_eq_acom.induct)  | 
| 47818 | 84  | 
apply (auto intro: le_trans simp: SKIP_le Assign_le Seq_le If_le While_le)  | 
| 
45623
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents:  
diff
changeset
 | 
85  | 
done  | 
| 
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents:  
diff
changeset
 | 
86  | 
next  | 
| 
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents:  
diff
changeset
 | 
87  | 
case goal4 thus ?case  | 
| 
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents:  
diff
changeset
 | 
88  | 
apply(induct x y rule: less_eq_acom.induct)  | 
| 
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents:  
diff
changeset
 | 
89  | 
apply (auto intro: le_antisym)  | 
| 
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents:  
diff
changeset
 | 
90  | 
done  | 
| 
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents:  
diff
changeset
 | 
91  | 
qed  | 
| 
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents:  
diff
changeset
 | 
92  | 
|
| 
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents:  
diff
changeset
 | 
93  | 
end  | 
| 
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents:  
diff
changeset
 | 
94  | 
|
| 45919 | 95  | 
fun sub\<^isub>1 :: "'a acom \<Rightarrow> 'a acom" where  | 
| 49095 | 96  | 
"sub\<^isub>1(C1;C2) = C1" |  | 
| 49344 | 97  | 
"sub\<^isub>1(IF b THEN {P1} C1 ELSE {P2} C2 {Q}) = C1" |
 | 
98  | 
"sub\<^isub>1({I} WHILE b DO {P} C {Q}) = C"
 | 
|
| 45903 | 99  | 
|
| 45919 | 100  | 
fun sub\<^isub>2 :: "'a acom \<Rightarrow> 'a acom" where  | 
| 49095 | 101  | 
"sub\<^isub>2(C1;C2) = C2" |  | 
| 49344 | 102  | 
"sub\<^isub>2(IF b THEN {P1} C1 ELSE {P2} C2 {Q}) = C2"
 | 
| 45903 | 103  | 
|
| 49095 | 104  | 
fun anno\<^isub>1 :: "'a acom \<Rightarrow> 'a" where  | 
| 49344 | 105  | 
"anno\<^isub>1(IF b THEN {P1} C1 ELSE {P2} C2 {Q}) = P1" |
 | 
106  | 
"anno\<^isub>1({I} WHILE b DO {P} C {Q}) = I"
 | 
|
| 49095 | 107  | 
|
108  | 
fun anno\<^isub>2 :: "'a acom \<Rightarrow> 'a" where  | 
|
| 49344 | 109  | 
"anno\<^isub>2(IF b THEN {P1} C1 ELSE {P2} C2 {Q}) = P2" |
 | 
110  | 
"anno\<^isub>2({I} WHILE b DO {P} C {Q}) = P"
 | 
|
| 49095 | 111  | 
|
| 45903 | 112  | 
|
| 49397 | 113  | 
fun Union_acom :: "com \<Rightarrow> 'a acom set \<Rightarrow> 'a set acom" where  | 
114  | 
"Union_acom com.SKIP M = (SKIP {post ` M})" |
 | 
|
115  | 
"Union_acom (x ::= a) M = (x ::= a {post ` M})" |
 | 
|
116  | 
"Union_acom (c1;c2) M =  | 
|
117  | 
Union_acom c1 (sub\<^isub>1 ` M); Union_acom c2 (sub\<^isub>2 ` M)" |  | 
|
118  | 
"Union_acom (IF b THEN c1 ELSE c2) M =  | 
|
119  | 
  IF b THEN {anno\<^isub>1 ` M} Union_acom c1 (sub\<^isub>1 ` M) ELSE {anno\<^isub>2 ` M} Union_acom c2 (sub\<^isub>2 ` M)
 | 
|
120  | 
  {post ` M}" |
 | 
|
121  | 
"Union_acom (WHILE b DO c) M =  | 
|
122  | 
 {anno\<^isub>1 ` M}
 | 
|
123  | 
 WHILE b DO {anno\<^isub>2 ` M} Union_acom c (sub\<^isub>1 ` M)
 | 
|
124  | 
 {post ` M}"
 | 
|
| 
45623
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents:  
diff
changeset
 | 
125  | 
|
| 49397 | 126  | 
interpretation  | 
127  | 
  Complete_Lattice "{C. strip C = c}" "map_acom Inter o (Union_acom c)" for c
 | 
|
| 
45623
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents:  
diff
changeset
 | 
128  | 
proof  | 
| 
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents:  
diff
changeset
 | 
129  | 
case goal1  | 
| 49397 | 130  | 
have "a:A \<Longrightarrow> map_acom Inter (Union_acom (strip a) A) \<le> a"  | 
| 
45623
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents:  
diff
changeset
 | 
131  | 
proof(induction a arbitrary: A)  | 
| 47818 | 132  | 
case Seq from Seq.prems show ?case by(force intro!: Seq.IH)  | 
| 
45623
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents:  
diff
changeset
 | 
133  | 
next  | 
| 45903 | 134  | 
case If from If.prems show ?case by(force intro!: If.IH)  | 
| 
45623
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents:  
diff
changeset
 | 
135  | 
next  | 
| 45903 | 136  | 
case While from While.prems show ?case by(force intro!: While.IH)  | 
137  | 
qed force+  | 
|
| 
45623
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents:  
diff
changeset
 | 
138  | 
with goal1 show ?case by auto  | 
| 
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents:  
diff
changeset
 | 
139  | 
next  | 
| 
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents:  
diff
changeset
 | 
140  | 
case goal2  | 
| 
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents:  
diff
changeset
 | 
141  | 
thus ?case  | 
| 
48759
 
ff570720ba1c
Improved complete lattice formalisation - no more index set.
 
nipkow 
parents: 
47818 
diff
changeset
 | 
142  | 
proof(simp, induction b arbitrary: c A)  | 
| 45903 | 143  | 
case SKIP thus ?case by (force simp:SKIP_le)  | 
144  | 
next  | 
|
145  | 
case Assign thus ?case by (force simp:Assign_le)  | 
|
| 
45623
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents:  
diff
changeset
 | 
146  | 
next  | 
| 
48759
 
ff570720ba1c
Improved complete lattice formalisation - no more index set.
 
nipkow 
parents: 
47818 
diff
changeset
 | 
147  | 
case Seq from Seq.prems show ?case by(force intro!: Seq.IH simp:Seq_le)  | 
| 
45623
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents:  
diff
changeset
 | 
148  | 
next  | 
| 45903 | 149  | 
case If from If.prems show ?case by (force simp: If_le intro!: If.IH)  | 
150  | 
next  | 
|
| 
48759
 
ff570720ba1c
Improved complete lattice formalisation - no more index set.
 
nipkow 
parents: 
47818 
diff
changeset
 | 
151  | 
case While from While.prems show ?case by(fastforce simp: While_le intro: While.IH)  | 
| 45903 | 152  | 
qed  | 
| 
45623
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents:  
diff
changeset
 | 
153  | 
next  | 
| 
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents:  
diff
changeset
 | 
154  | 
case goal3  | 
| 49397 | 155  | 
have "strip(Union_acom c A) = c"  | 
| 
48759
 
ff570720ba1c
Improved complete lattice formalisation - no more index set.
 
nipkow 
parents: 
47818 
diff
changeset
 | 
156  | 
proof(induction c arbitrary: A)  | 
| 
 
ff570720ba1c
Improved complete lattice formalisation - no more index set.
 
nipkow 
parents: 
47818 
diff
changeset
 | 
157  | 
case Seq from Seq.prems show ?case by (fastforce simp: strip_eq_Seq subset_iff intro!: Seq.IH)  | 
| 
45623
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents:  
diff
changeset
 | 
158  | 
next  | 
| 
48759
 
ff570720ba1c
Improved complete lattice formalisation - no more index set.
 
nipkow 
parents: 
47818 
diff
changeset
 | 
159  | 
case If from If.prems show ?case by (fastforce intro!: If.IH simp: strip_eq_If)  | 
| 
45623
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents:  
diff
changeset
 | 
160  | 
next  | 
| 
48759
 
ff570720ba1c
Improved complete lattice formalisation - no more index set.
 
nipkow 
parents: 
47818 
diff
changeset
 | 
161  | 
case While from While.prems show ?case by(fastforce intro: While.IH simp: strip_eq_While)  | 
| 
45623
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents:  
diff
changeset
 | 
162  | 
qed auto  | 
| 45903 | 163  | 
thus ?case by auto  | 
| 
45623
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents:  
diff
changeset
 | 
164  | 
qed  | 
| 
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents:  
diff
changeset
 | 
165  | 
|
| 
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents:  
diff
changeset
 | 
166  | 
lemma le_post: "c \<le> d \<Longrightarrow> post c \<le> post d"  | 
| 
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents:  
diff
changeset
 | 
167  | 
by(induction c d rule: less_eq_acom.induct) auto  | 
| 
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents:  
diff
changeset
 | 
168  | 
|
| 49487 | 169  | 
|
| 
45623
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents:  
diff
changeset
 | 
170  | 
subsubsection "Collecting semantics"  | 
| 
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents:  
diff
changeset
 | 
171  | 
|
| 51389 | 172  | 
interpretation Step  | 
173  | 
where step_assign = "\<lambda>x e S. {s(x := aval e s) |s. s : S}"
 | 
|
174  | 
and step_cond = "\<lambda>b S. {s:S. bval b s}"
 | 
|
175  | 
defines step is Step  | 
|
176  | 
.  | 
|
| 
45623
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents:  
diff
changeset
 | 
177  | 
|
| 46070 | 178  | 
definition CS :: "com \<Rightarrow> state set acom" where  | 
| 
48759
 
ff570720ba1c
Improved complete lattice formalisation - no more index set.
 
nipkow 
parents: 
47818 
diff
changeset
 | 
179  | 
"CS c = lfp c (step UNIV)"  | 
| 
45623
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents:  
diff
changeset
 | 
180  | 
|
| 46334 | 181  | 
lemma mono2_step: "c1 \<le> c2 \<Longrightarrow> S1 \<subseteq> S2 \<Longrightarrow> step S1 c1 \<le> step S2 c2"  | 
182  | 
proof(induction c1 c2 arbitrary: S1 S2 rule: less_eq_acom.induct)  | 
|
| 
45623
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents:  
diff
changeset
 | 
183  | 
case 2 thus ?case by fastforce  | 
| 
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents:  
diff
changeset
 | 
184  | 
next  | 
| 
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents:  
diff
changeset
 | 
185  | 
case 3 thus ?case by(simp add: le_post)  | 
| 
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents:  
diff
changeset
 | 
186  | 
next  | 
| 
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents:  
diff
changeset
 | 
187  | 
case 4 thus ?case by(simp add: subset_iff)(metis le_post set_mp)+  | 
| 
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents:  
diff
changeset
 | 
188  | 
next  | 
| 
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents:  
diff
changeset
 | 
189  | 
case 5 thus ?case by(simp add: subset_iff) (metis le_post set_mp)  | 
| 
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents:  
diff
changeset
 | 
190  | 
qed auto  | 
| 
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents:  
diff
changeset
 | 
191  | 
|
| 
45655
 
a49f9428aba4
simplified Collecting1 and renamed: step -> step', step_cs -> step
 
nipkow 
parents: 
45623 
diff
changeset
 | 
192  | 
lemma mono_step: "mono (step S)"  | 
| 46334 | 193  | 
by(blast intro: monoI mono2_step)  | 
| 
45623
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents:  
diff
changeset
 | 
194  | 
|
| 49095 | 195  | 
lemma strip_step: "strip(step S C) = strip C"  | 
196  | 
by (induction C arbitrary: S) auto  | 
|
| 
45623
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents:  
diff
changeset
 | 
197  | 
|
| 
48759
 
ff570720ba1c
Improved complete lattice formalisation - no more index set.
 
nipkow 
parents: 
47818 
diff
changeset
 | 
198  | 
lemma lfp_cs_unfold: "lfp c (step S) = step S (lfp c (step S))"  | 
| 45903 | 199  | 
apply(rule lfp_unfold[OF _ mono_step])  | 
200  | 
apply(simp add: strip_step)  | 
|
201  | 
done  | 
|
| 
45623
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents:  
diff
changeset
 | 
202  | 
|
| 46070 | 203  | 
lemma CS_unfold: "CS c = step UNIV (CS c)"  | 
| 
45623
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents:  
diff
changeset
 | 
204  | 
by (metis CS_def lfp_cs_unfold)  | 
| 
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents:  
diff
changeset
 | 
205  | 
|
| 46070 | 206  | 
lemma strip_CS[simp]: "strip(CS c) = c"  | 
| 45903 | 207  | 
by(simp add: CS_def index_lfp[simplified])  | 
| 
45623
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents:  
diff
changeset
 | 
208  | 
|
| 49487 | 209  | 
|
210  | 
subsubsection "Relation to big-step semantics"  | 
|
211  | 
||
212  | 
lemma post_Union_acom: "\<forall> c' \<in> M. strip c' = c \<Longrightarrow> post (Union_acom c M) = post ` M"  | 
|
213  | 
proof(induction c arbitrary: M)  | 
|
214  | 
case (Seq c1 c2)  | 
|
215  | 
have "post ` M = post ` sub\<^isub>2 ` M" using Seq.prems by (force simp: strip_eq_Seq)  | 
|
216  | 
moreover have "\<forall> c' \<in> sub\<^isub>2 ` M. strip c' = c2" using Seq.prems by (auto simp: strip_eq_Seq)  | 
|
217  | 
ultimately show ?case using Seq.IH(2)[of "sub\<^isub>2 ` M"] by simp  | 
|
218  | 
qed simp_all  | 
|
219  | 
||
220  | 
||
221  | 
lemma post_lfp: "post(lfp c f) = (\<Inter>{post C|C. strip C = c \<and> f C \<le> C})"
 | 
|
222  | 
by(auto simp add: lfp_def post_Union_acom)  | 
|
223  | 
||
224  | 
lemma big_step_post_step:  | 
|
225  | 
"\<lbrakk> (c, s) \<Rightarrow> t; strip C = c; s \<in> S; step S C \<le> C \<rbrakk> \<Longrightarrow> t \<in> post C"  | 
|
226  | 
proof(induction arbitrary: C S rule: big_step_induct)  | 
|
227  | 
case Skip thus ?case by(auto simp: strip_eq_SKIP)  | 
|
228  | 
next  | 
|
229  | 
case Assign thus ?case by(fastforce simp: strip_eq_Assign)  | 
|
230  | 
next  | 
|
231  | 
case Seq thus ?case by(fastforce simp: strip_eq_Seq)  | 
|
232  | 
next  | 
|
233  | 
case IfTrue thus ?case apply(auto simp: strip_eq_If)  | 
|
234  | 
by (metis (lifting) mem_Collect_eq set_mp)  | 
|
235  | 
next  | 
|
236  | 
case IfFalse thus ?case apply(auto simp: strip_eq_If)  | 
|
237  | 
by (metis (lifting) mem_Collect_eq set_mp)  | 
|
238  | 
next  | 
|
239  | 
case (WhileTrue b s1 c' s2 s3)  | 
|
240  | 
  from WhileTrue.prems(1) obtain I P C' Q where "C = {I} WHILE b DO {P} C' {Q}" "strip C' = c'"
 | 
|
241  | 
by(auto simp: strip_eq_While)  | 
|
242  | 
from WhileTrue.prems(3) `C = _`  | 
|
243  | 
  have "step P C' \<le> C'"  "{s \<in> I. bval b s} \<le> P"  "S \<le> I"  "step (post C') C \<le> C"  by auto
 | 
|
244  | 
  have "step {s \<in> I. bval b s} C' \<le> C'"
 | 
|
245  | 
    by (rule order_trans[OF mono2_step[OF order_refl `{s \<in> I. bval b s} \<le> P`] `step P C' \<le> C'`])
 | 
|
246  | 
  have "s1: {s:I. bval b s}" using `s1 \<in> S` `S \<subseteq> I` `bval b s1` by auto
 | 
|
247  | 
  note s2_in_post_C' = WhileTrue.IH(1)[OF `strip C' = c'` this `step {s \<in> I. bval b s} C' \<le> C'`]
 | 
|
248  | 
from WhileTrue.IH(2)[OF WhileTrue.prems(1) s2_in_post_C' `step (post C') C \<le> C`]  | 
|
249  | 
show ?case .  | 
|
250  | 
next  | 
|
251  | 
case (WhileFalse b s1 c') thus ?case by (force simp: strip_eq_While)  | 
|
252  | 
qed  | 
|
253  | 
||
254  | 
lemma big_step_lfp: "\<lbrakk> (c,s) \<Rightarrow> t; s \<in> S \<rbrakk> \<Longrightarrow> t \<in> post(lfp c (step S))"  | 
|
255  | 
by(auto simp add: post_lfp intro: big_step_post_step)  | 
|
256  | 
||
257  | 
lemma big_step_CS: "(c,s) \<Rightarrow> t \<Longrightarrow> t : post(CS c)"  | 
|
258  | 
by(simp add: CS_def big_step_lfp)  | 
|
259  | 
||
| 
45623
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents:  
diff
changeset
 | 
260  | 
end  |