| author | haftmann | 
| Fri, 14 Jun 2019 08:34:27 +0000 | |
| changeset 70332 | 315489d836d8 | 
| parent 68451 | c34aa23a1fb6 | 
| permissions | -rw-r--r-- | 
| 52072 | 1  | 
theory Fold imports Sem_Equiv Vars begin  | 
| 44070 | 2  | 
|
| 44850 | 3  | 
subsection "Simple folding of arithmetic expressions"  | 
| 44070 | 4  | 
|
| 45134 | 5  | 
type_synonym  | 
| 45212 | 6  | 
tab = "vname \<Rightarrow> val option"  | 
| 44070 | 7  | 
|
| 
51514
 
9bed72e55475
simp_const -> afold; bfold -> fold'; bsimp_const -> bfold
 
kleing 
parents: 
51506 
diff
changeset
 | 
8  | 
fun afold :: "aexp \<Rightarrow> tab \<Rightarrow> aexp" where  | 
| 
 
9bed72e55475
simp_const -> afold; bfold -> fold'; bsimp_const -> bfold
 
kleing 
parents: 
51506 
diff
changeset
 | 
9  | 
"afold (N n) _ = N n" |  | 
| 
 
9bed72e55475
simp_const -> afold; bfold -> fold'; bsimp_const -> bfold
 
kleing 
parents: 
51506 
diff
changeset
 | 
10  | 
"afold (V x) t = (case t x of None \<Rightarrow> V x | Some k \<Rightarrow> N k)" |  | 
| 
 
9bed72e55475
simp_const -> afold; bfold -> fold'; bsimp_const -> bfold
 
kleing 
parents: 
51506 
diff
changeset
 | 
11  | 
"afold (Plus e1 e2) t = (case (afold e1 t, afold e2 t) of  | 
| 
48909
 
b2dea007e55e
remove Hoare dependency from Fold.thy
 
Gerwin Klein <gerwin.klein@nicta.com.au> 
parents: 
47818 
diff
changeset
 | 
12  | 
(N n1, N n2) \<Rightarrow> N(n1+n2) | (e1',e2') \<Rightarrow> Plus e1' e2')"  | 
| 44070 | 13  | 
|
| 67613 | 14  | 
definition "approx t s \<longleftrightarrow> (\<forall>x k. t x = Some k \<longrightarrow> s x = k)"  | 
| 44070 | 15  | 
|
| 
51514
 
9bed72e55475
simp_const -> afold; bfold -> fold'; bsimp_const -> bfold
 
kleing 
parents: 
51506 
diff
changeset
 | 
16  | 
theorem aval_afold[simp]:  | 
| 44070 | 17  | 
assumes "approx t s"  | 
| 
51514
 
9bed72e55475
simp_const -> afold; bfold -> fold'; bsimp_const -> bfold
 
kleing 
parents: 
51506 
diff
changeset
 | 
18  | 
shows "aval (afold a t) s = aval a s"  | 
| 
48909
 
b2dea007e55e
remove Hoare dependency from Fold.thy
 
Gerwin Klein <gerwin.klein@nicta.com.au> 
parents: 
47818 
diff
changeset
 | 
19  | 
using assms  | 
| 44070 | 20  | 
by (induct a) (auto simp: approx_def split: aexp.split option.split)  | 
21  | 
||
| 
51514
 
9bed72e55475
simp_const -> afold; bfold -> fold'; bsimp_const -> bfold
 
kleing 
parents: 
51506 
diff
changeset
 | 
22  | 
theorem aval_afold_N:  | 
| 44070 | 23  | 
assumes "approx t s"  | 
| 
51514
 
9bed72e55475
simp_const -> afold; bfold -> fold'; bsimp_const -> bfold
 
kleing 
parents: 
51506 
diff
changeset
 | 
24  | 
shows "afold a t = N n \<Longrightarrow> aval a s = n"  | 
| 51573 | 25  | 
by (metis assms aval.simps(1) aval_afold)  | 
| 44070 | 26  | 
|
27  | 
definition  | 
|
28  | 
"merge t1 t2 = (\<lambda>m. if t1 m = t2 m then t1 m else None)"  | 
|
29  | 
||
30  | 
primrec "defs" :: "com \<Rightarrow> tab \<Rightarrow> tab" where  | 
|
31  | 
"defs SKIP t = t" |  | 
|
32  | 
"defs (x ::= a) t =  | 
|
| 
51514
 
9bed72e55475
simp_const -> afold; bfold -> fold'; bsimp_const -> bfold
 
kleing 
parents: 
51506 
diff
changeset
 | 
33  | 
(case afold a t of N k \<Rightarrow> t(x \<mapsto> k) | _ \<Rightarrow> t(x:=None))" |  | 
| 
52046
 
bc01725d7918
replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
 
nipkow 
parents: 
51573 
diff
changeset
 | 
34  | 
"defs (c1;;c2) t = (defs c2 o defs c1) t" |  | 
| 44070 | 35  | 
"defs (IF b THEN c1 ELSE c2) t = merge (defs c1 t) (defs c2 t)" |  | 
| 52072 | 36  | 
"defs (WHILE b DO c) t = t |` (-lvars c)"  | 
| 44070 | 37  | 
|
38  | 
primrec fold where  | 
|
39  | 
"fold SKIP _ = SKIP" |  | 
|
| 
51514
 
9bed72e55475
simp_const -> afold; bfold -> fold'; bsimp_const -> bfold
 
kleing 
parents: 
51506 
diff
changeset
 | 
40  | 
"fold (x ::= a) t = (x ::= (afold a t))" |  | 
| 
52046
 
bc01725d7918
replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
 
nipkow 
parents: 
51573 
diff
changeset
 | 
41  | 
"fold (c1;;c2) t = (fold c1 t;; fold c2 (defs c1 t))" |  | 
| 44070 | 42  | 
"fold (IF b THEN c1 ELSE c2) t = IF b THEN fold c1 t ELSE fold c2 t" |  | 
| 52072 | 43  | 
"fold (WHILE b DO c) t = WHILE b DO fold c (t |` (-lvars c))"  | 
| 44070 | 44  | 
|
45  | 
lemma approx_merge:  | 
|
46  | 
"approx t1 s \<or> approx t2 s \<Longrightarrow> approx (merge t1 t2) s"  | 
|
| 
44890
 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 
nipkow 
parents: 
44850 
diff
changeset
 | 
47  | 
by (fastforce simp: merge_def approx_def)  | 
| 44070 | 48  | 
|
49  | 
lemma approx_map_le:  | 
|
50  | 
"approx t2 s \<Longrightarrow> t1 \<subseteq>\<^sub>m t2 \<Longrightarrow> approx t1 s"  | 
|
51  | 
by (clarsimp simp: approx_def map_le_def dom_def)  | 
|
52  | 
||
53  | 
lemma restrict_map_le [intro!, simp]: "t |` S \<subseteq>\<^sub>m t"  | 
|
54  | 
by (clarsimp simp: restrict_map_def map_le_def)  | 
|
55  | 
||
56  | 
lemma merge_restrict:  | 
|
57  | 
assumes "t1 |` S = t |` S"  | 
|
58  | 
assumes "t2 |` S = t |` S"  | 
|
59  | 
shows "merge t1 t2 |` S = t |` S"  | 
|
60  | 
proof -  | 
|
61  | 
from assms  | 
|
| 
48909
 
b2dea007e55e
remove Hoare dependency from Fold.thy
 
Gerwin Klein <gerwin.klein@nicta.com.au> 
parents: 
47818 
diff
changeset
 | 
62  | 
have "\<forall>x. (t1 |` S) x = (t |` S) x"  | 
| 44070 | 63  | 
and "\<forall>x. (t2 |` S) x = (t |` S) x" by auto  | 
64  | 
thus ?thesis  | 
|
| 
48909
 
b2dea007e55e
remove Hoare dependency from Fold.thy
 
Gerwin Klein <gerwin.klein@nicta.com.au> 
parents: 
47818 
diff
changeset
 | 
65  | 
by (auto simp: merge_def restrict_map_def  | 
| 51506 | 66  | 
split: if_splits)  | 
| 44070 | 67  | 
qed  | 
68  | 
||
69  | 
||
70  | 
lemma defs_restrict:  | 
|
| 52072 | 71  | 
"defs c t |` (- lvars c) = t |` (- lvars c)"  | 
| 45015 | 72  | 
proof (induction c arbitrary: t)  | 
| 47818 | 73  | 
case (Seq c1 c2)  | 
| 52072 | 74  | 
hence "defs c1 t |` (- lvars c1) = t |` (- lvars c1)"  | 
| 44070 | 75  | 
by simp  | 
| 52072 | 76  | 
hence "defs c1 t |` (- lvars c1) |` (-lvars c2) =  | 
77  | 
t |` (- lvars c1) |` (-lvars c2)" by simp  | 
|
| 44070 | 78  | 
moreover  | 
| 47818 | 79  | 
from Seq  | 
| 52072 | 80  | 
have "defs c2 (defs c1 t) |` (- lvars c2) =  | 
81  | 
defs c1 t |` (- lvars c2)"  | 
|
| 44070 | 82  | 
by simp  | 
| 52072 | 83  | 
hence "defs c2 (defs c1 t) |` (- lvars c2) |` (- lvars c1) =  | 
84  | 
defs c1 t |` (- lvars c2) |` (- lvars c1)"  | 
|
| 44070 | 85  | 
by simp  | 
86  | 
ultimately  | 
|
87  | 
show ?case by (clarsimp simp: Int_commute)  | 
|
88  | 
next  | 
|
89  | 
case (If b c1 c2)  | 
|
| 52072 | 90  | 
hence "defs c1 t |` (- lvars c1) = t |` (- lvars c1)" by simp  | 
91  | 
hence "defs c1 t |` (- lvars c1) |` (-lvars c2) =  | 
|
92  | 
t |` (- lvars c1) |` (-lvars c2)" by simp  | 
|
| 44070 | 93  | 
moreover  | 
94  | 
from If  | 
|
| 52072 | 95  | 
have "defs c2 t |` (- lvars c2) = t |` (- lvars c2)" by simp  | 
96  | 
hence "defs c2 t |` (- lvars c2) |` (-lvars c1) =  | 
|
97  | 
t |` (- lvars c2) |` (-lvars c1)" by simp  | 
|
| 44070 | 98  | 
ultimately  | 
99  | 
show ?case by (auto simp: Int_commute intro: merge_restrict)  | 
|
100  | 
qed (auto split: aexp.split)  | 
|
101  | 
||
102  | 
||
103  | 
lemma big_step_pres_approx:  | 
|
104  | 
"(c,s) \<Rightarrow> s' \<Longrightarrow> approx t s \<Longrightarrow> approx (defs c t) s'"  | 
|
| 45015 | 105  | 
proof (induction arbitrary: t rule: big_step_induct)  | 
| 44070 | 106  | 
case Skip thus ?case by simp  | 
107  | 
next  | 
|
108  | 
case Assign  | 
|
109  | 
thus ?case  | 
|
| 
51514
 
9bed72e55475
simp_const -> afold; bfold -> fold'; bsimp_const -> bfold
 
kleing 
parents: 
51506 
diff
changeset
 | 
110  | 
by (clarsimp simp: aval_afold_N approx_def split: aexp.split)  | 
| 44070 | 111  | 
next  | 
| 47818 | 112  | 
case (Seq c1 s1 s2 c2 s3)  | 
113  | 
have "approx (defs c1 t) s2" by (rule Seq.IH(1)[OF Seq.prems])  | 
|
114  | 
hence "approx (defs c2 (defs c1 t)) s3" by (rule Seq.IH(2))  | 
|
| 44070 | 115  | 
thus ?case by simp  | 
116  | 
next  | 
|
117  | 
case (IfTrue b s c1 s')  | 
|
118  | 
hence "approx (defs c1 t) s'" by simp  | 
|
119  | 
thus ?case by (simp add: approx_merge)  | 
|
120  | 
next  | 
|
121  | 
case (IfFalse b s c2 s')  | 
|
122  | 
hence "approx (defs c2 t) s'" by simp  | 
|
123  | 
thus ?case by (simp add: approx_merge)  | 
|
124  | 
next  | 
|
125  | 
case WhileFalse  | 
|
126  | 
thus ?case by (simp add: approx_def restrict_map_def)  | 
|
127  | 
next  | 
|
128  | 
case (WhileTrue b s1 c s2 s3)  | 
|
129  | 
hence "approx (defs c t) s2" by simp  | 
|
130  | 
with WhileTrue  | 
|
| 52072 | 131  | 
have "approx (defs c t |` (-lvars c)) s3" by simp  | 
| 44070 | 132  | 
thus ?case by (simp add: defs_restrict)  | 
133  | 
qed  | 
|
134  | 
||
135  | 
||
136  | 
lemma big_step_pres_approx_restrict:  | 
|
| 52072 | 137  | 
"(c,s) \<Rightarrow> s' \<Longrightarrow> approx (t |` (-lvars c)) s \<Longrightarrow> approx (t |` (-lvars c)) s'"  | 
| 45015 | 138  | 
proof (induction arbitrary: t rule: big_step_induct)  | 
| 44070 | 139  | 
case Assign  | 
140  | 
thus ?case by (clarsimp simp: approx_def)  | 
|
141  | 
next  | 
|
| 47818 | 142  | 
case (Seq c1 s1 s2 c2 s3)  | 
| 52072 | 143  | 
hence "approx (t |` (-lvars c2) |` (-lvars c1)) s1"  | 
| 44070 | 144  | 
by (simp add: Int_commute)  | 
| 52072 | 145  | 
hence "approx (t |` (-lvars c2) |` (-lvars c1)) s2"  | 
| 47818 | 146  | 
by (rule Seq)  | 
| 52072 | 147  | 
hence "approx (t |` (-lvars c1) |` (-lvars c2)) s2"  | 
| 44070 | 148  | 
by (simp add: Int_commute)  | 
| 52072 | 149  | 
hence "approx (t |` (-lvars c1) |` (-lvars c2)) s3"  | 
| 47818 | 150  | 
by (rule Seq)  | 
| 44070 | 151  | 
thus ?case by simp  | 
152  | 
next  | 
|
153  | 
case (IfTrue b s c1 s' c2)  | 
|
| 52072 | 154  | 
hence "approx (t |` (-lvars c2) |` (-lvars c1)) s"  | 
| 44070 | 155  | 
by (simp add: Int_commute)  | 
| 52072 | 156  | 
hence "approx (t |` (-lvars c2) |` (-lvars c1)) s'"  | 
| 44070 | 157  | 
by (rule IfTrue)  | 
| 
48909
 
b2dea007e55e
remove Hoare dependency from Fold.thy
 
Gerwin Klein <gerwin.klein@nicta.com.au> 
parents: 
47818 
diff
changeset
 | 
158  | 
thus ?case by (simp add: Int_commute)  | 
| 44070 | 159  | 
next  | 
160  | 
case (IfFalse b s c2 s' c1)  | 
|
| 52072 | 161  | 
hence "approx (t |` (-lvars c1) |` (-lvars c2)) s"  | 
| 44070 | 162  | 
by simp  | 
| 52072 | 163  | 
hence "approx (t |` (-lvars c1) |` (-lvars c2)) s'"  | 
| 44070 | 164  | 
by (rule IfFalse)  | 
165  | 
thus ?case by simp  | 
|
166  | 
qed auto  | 
|
167  | 
||
168  | 
||
169  | 
declare assign_simp [simp]  | 
|
170  | 
||
171  | 
lemma approx_eq:  | 
|
172  | 
"approx t \<Turnstile> c \<sim> fold c t"  | 
|
| 45015 | 173  | 
proof (induction c arbitrary: t)  | 
| 44070 | 174  | 
case SKIP show ?case by simp  | 
175  | 
next  | 
|
176  | 
case Assign  | 
|
177  | 
show ?case by (simp add: equiv_up_to_def)  | 
|
178  | 
next  | 
|
| 
48909
 
b2dea007e55e
remove Hoare dependency from Fold.thy
 
Gerwin Klein <gerwin.klein@nicta.com.au> 
parents: 
47818 
diff
changeset
 | 
179  | 
case Seq  | 
| 
 
b2dea007e55e
remove Hoare dependency from Fold.thy
 
Gerwin Klein <gerwin.klein@nicta.com.au> 
parents: 
47818 
diff
changeset
 | 
180  | 
thus ?case by (auto intro!: equiv_up_to_seq big_step_pres_approx)  | 
| 44070 | 181  | 
next  | 
182  | 
case If  | 
|
183  | 
thus ?case by (auto intro!: equiv_up_to_if_weak)  | 
|
184  | 
next  | 
|
185  | 
case (While b c)  | 
|
| 52072 | 186  | 
hence "approx (t |` (- lvars c)) \<Turnstile>  | 
187  | 
WHILE b DO c \<sim> WHILE b DO fold c (t |` (- lvars c))"  | 
|
| 
48909
 
b2dea007e55e
remove Hoare dependency from Fold.thy
 
Gerwin Klein <gerwin.klein@nicta.com.au> 
parents: 
47818 
diff
changeset
 | 
188  | 
by (auto intro: equiv_up_to_while_weak big_step_pres_approx_restrict)  | 
| 
 
b2dea007e55e
remove Hoare dependency from Fold.thy
 
Gerwin Klein <gerwin.klein@nicta.com.au> 
parents: 
47818 
diff
changeset
 | 
189  | 
thus ?case  | 
| 44070 | 190  | 
by (auto intro: equiv_up_to_weaken approx_map_le)  | 
191  | 
qed  | 
|
| 
48909
 
b2dea007e55e
remove Hoare dependency from Fold.thy
 
Gerwin Klein <gerwin.klein@nicta.com.au> 
parents: 
47818 
diff
changeset
 | 
192  | 
|
| 44070 | 193  | 
|
| 
48909
 
b2dea007e55e
remove Hoare dependency from Fold.thy
 
Gerwin Klein <gerwin.klein@nicta.com.au> 
parents: 
47818 
diff
changeset
 | 
194  | 
lemma approx_empty [simp]:  | 
| 68451 | 195  | 
"approx Map.empty = (\<lambda>_. True)"  | 
| 51506 | 196  | 
by (auto simp: approx_def)  | 
| 44070 | 197  | 
|
198  | 
||
199  | 
theorem constant_folding_equiv:  | 
|
| 68451 | 200  | 
"fold c Map.empty \<sim> c"  | 
201  | 
using approx_eq [of Map.empty c]  | 
|
| 52825 | 202  | 
by (simp add: equiv_up_to_True sim_sym)  | 
| 44070 | 203  | 
|
204  | 
||
205  | 
end  |