| author | wenzelm | 
| Tue, 04 Nov 2014 17:37:15 +0100 | |
| changeset 58896 | 5a2f475e2ded | 
| parent 58310 | 91ea607a34d8 | 
| child 60170 | 031ec3d34d31 | 
| permissions | -rw-r--r-- | 
| 43141 | 1  | 
theory BExp imports AExp begin  | 
2  | 
||
3  | 
subsection "Boolean Expressions"  | 
|
4  | 
||
| 58310 | 5  | 
datatype bexp = Bc bool | Not bexp | And bexp bexp | Less aexp aexp  | 
| 43141 | 6  | 
|
| 49191 | 7  | 
text_raw{*\snip{BExpbvaldef}{1}{2}{% *}
 | 
| 43141 | 8  | 
fun bval :: "bexp \<Rightarrow> state \<Rightarrow> bool" where  | 
| 45216 | 9  | 
"bval (Bc v) s = v" |  | 
| 43141 | 10  | 
"bval (Not b) s = (\<not> bval b s)" |  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
49920 
diff
changeset
 | 
11  | 
"bval (And b\<^sub>1 b\<^sub>2) s = (bval b\<^sub>1 s \<and> bval b\<^sub>2 s)" |  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
49920 
diff
changeset
 | 
12  | 
"bval (Less a\<^sub>1 a\<^sub>2) s = (aval a\<^sub>1 s < aval a\<^sub>2 s)"  | 
| 49191 | 13  | 
text_raw{*}%endsnip*}
 | 
| 43141 | 14  | 
|
15  | 
value "bval (Less (V ''x'') (Plus (N 3) (V ''y'')))  | 
|
| 44036 | 16  | 
<''x'' := 3, ''y'' := 1>"  | 
| 43141 | 17  | 
|
18  | 
||
| 45255 | 19  | 
text{* To improve automation: *}
 | 
| 43141 | 20  | 
|
| 45255 | 21  | 
lemma bval_And_if[simp]:  | 
22  | 
"bval (And b1 b2) s = (if bval b1 s then bval b2 s else False)"  | 
|
23  | 
by(simp)  | 
|
24  | 
||
25  | 
declare bval.simps(3)[simp del] --"remove the original eqn"  | 
|
26  | 
||
| 43141 | 27  | 
|
| 45255 | 28  | 
subsection "Constant Folding"  | 
29  | 
||
30  | 
text{* Optimizing constructors: *}
 | 
|
31  | 
||
| 49191 | 32  | 
text_raw{*\snip{BExplessdef}{0}{2}{% *}
 | 
| 43141 | 33  | 
fun less :: "aexp \<Rightarrow> aexp \<Rightarrow> bexp" where  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
49920 
diff
changeset
 | 
34  | 
"less (N n\<^sub>1) (N n\<^sub>2) = Bc(n\<^sub>1 < n\<^sub>2)" |  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
49920 
diff
changeset
 | 
35  | 
"less a\<^sub>1 a\<^sub>2 = Less a\<^sub>1 a\<^sub>2"  | 
| 49191 | 36  | 
text_raw{*}%endsnip*}
 | 
| 43141 | 37  | 
|
38  | 
lemma [simp]: "bval (less a1 a2) s = (aval a1 s < aval a2 s)"  | 
|
| 45015 | 39  | 
apply(induction a1 a2 rule: less.induct)  | 
| 43141 | 40  | 
apply simp_all  | 
41  | 
done  | 
|
42  | 
||
| 49191 | 43  | 
text_raw{*\snip{BExpanddef}{2}{2}{% *}
 | 
| 43141 | 44  | 
fun "and" :: "bexp \<Rightarrow> bexp \<Rightarrow> bexp" where  | 
| 45200 | 45  | 
"and (Bc True) b = b" |  | 
46  | 
"and b (Bc True) = b" |  | 
|
47  | 
"and (Bc False) b = Bc False" |  | 
|
48  | 
"and b (Bc False) = Bc False" |  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
49920 
diff
changeset
 | 
49  | 
"and b\<^sub>1 b\<^sub>2 = And b\<^sub>1 b\<^sub>2"  | 
| 49191 | 50  | 
text_raw{*}%endsnip*}
 | 
| 43141 | 51  | 
|
52  | 
lemma bval_and[simp]: "bval (and b1 b2) s = (bval b1 s \<and> bval b2 s)"  | 
|
| 45015 | 53  | 
apply(induction b1 b2 rule: and.induct)  | 
| 43141 | 54  | 
apply simp_all  | 
55  | 
done  | 
|
56  | 
||
| 49191 | 57  | 
text_raw{*\snip{BExpnotdef}{2}{2}{% *}
 | 
| 43141 | 58  | 
fun not :: "bexp \<Rightarrow> bexp" where  | 
| 45200 | 59  | 
"not (Bc True) = Bc False" |  | 
60  | 
"not (Bc False) = Bc True" |  | 
|
| 43141 | 61  | 
"not b = Not b"  | 
| 49191 | 62  | 
text_raw{*}%endsnip*}
 | 
| 43141 | 63  | 
|
| 45255 | 64  | 
lemma bval_not[simp]: "bval (not b) s = (\<not> bval b s)"  | 
| 45015 | 65  | 
apply(induction b rule: not.induct)  | 
| 43141 | 66  | 
apply simp_all  | 
67  | 
done  | 
|
68  | 
||
69  | 
text{* Now the overall optimizer: *}
 | 
|
70  | 
||
| 49191 | 71  | 
text_raw{*\snip{BExpbsimpdef}{0}{2}{% *}
 | 
| 43141 | 72  | 
fun bsimp :: "bexp \<Rightarrow> bexp" where  | 
| 45256 | 73  | 
"bsimp (Bc v) = Bc v" |  | 
74  | 
"bsimp (Not b) = not(bsimp b)" |  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
49920 
diff
changeset
 | 
75  | 
"bsimp (And b\<^sub>1 b\<^sub>2) = and (bsimp b\<^sub>1) (bsimp b\<^sub>2)" |  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
49920 
diff
changeset
 | 
76  | 
"bsimp (Less a\<^sub>1 a\<^sub>2) = less (asimp a\<^sub>1) (asimp a\<^sub>2)"  | 
| 49191 | 77  | 
text_raw{*}%endsnip*}
 | 
| 43141 | 78  | 
|
79  | 
value "bsimp (And (Less (N 0) (N 1)) b)"  | 
|
80  | 
||
| 49920 | 81  | 
value "bsimp (And (Less (N 1) (N 0)) (Bc True))"  | 
| 43141 | 82  | 
|
83  | 
theorem "bval (bsimp b) s = bval b s"  | 
|
| 45015 | 84  | 
apply(induction b)  | 
| 43141 | 85  | 
apply simp_all  | 
86  | 
done  | 
|
87  | 
||
88  | 
end  |