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