author | wenzelm |
Sat, 13 Aug 2022 18:06:30 +0200 | |
changeset 75848 | 9e4c0aaa30aa |
parent 68776 | 403dd13cf6e9 |
permissions | -rw-r--r-- |
68776
403dd13cf6e9
avoid session qualification because no tex is generated when used;
nipkow
parents:
67406
diff
changeset
|
1 |
subsection "Boolean Expressions" |
43141 | 2 |
|
68776
403dd13cf6e9
avoid session qualification because no tex is generated when used;
nipkow
parents:
67406
diff
changeset
|
3 |
theory BExp imports AExp begin |
43141 | 4 |
|
58310 | 5 |
datatype bexp = Bc bool | Not bexp | And bexp bexp | Less aexp aexp |
43141 | 6 |
|
67406 | 7 |
text_raw\<open>\snip{BExpbvaldef}{1}{2}{%\<close> |
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)" |
67406 | 13 |
text_raw\<open>}%endsnip\<close> |
43141 | 14 |
|
15 |
value "bval (Less (V ''x'') (Plus (N 3) (V ''y''))) |
|
44036 | 16 |
<''x'' := 3, ''y'' := 1>" |
43141 | 17 |
|
18 |
||
45255 | 19 |
subsection "Constant Folding" |
20 |
||
67406 | 21 |
text\<open>Optimizing constructors:\<close> |
45255 | 22 |
|
67406 | 23 |
text_raw\<open>\snip{BExplessdef}{0}{2}{%\<close> |
43141 | 24 |
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
|
25 |
"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
|
26 |
"less a\<^sub>1 a\<^sub>2 = Less a\<^sub>1 a\<^sub>2" |
67406 | 27 |
text_raw\<open>}%endsnip\<close> |
43141 | 28 |
|
29 |
lemma [simp]: "bval (less a1 a2) s = (aval a1 s < aval a2 s)" |
|
45015 | 30 |
apply(induction a1 a2 rule: less.induct) |
43141 | 31 |
apply simp_all |
32 |
done |
|
33 |
||
67406 | 34 |
text_raw\<open>\snip{BExpanddef}{2}{2}{%\<close> |
43141 | 35 |
fun "and" :: "bexp \<Rightarrow> bexp \<Rightarrow> bexp" where |
45200 | 36 |
"and (Bc True) b = b" | |
37 |
"and b (Bc True) = b" | |
|
38 |
"and (Bc False) b = Bc False" | |
|
39 |
"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
|
40 |
"and b\<^sub>1 b\<^sub>2 = And b\<^sub>1 b\<^sub>2" |
67406 | 41 |
text_raw\<open>}%endsnip\<close> |
43141 | 42 |
|
43 |
lemma bval_and[simp]: "bval (and b1 b2) s = (bval b1 s \<and> bval b2 s)" |
|
45015 | 44 |
apply(induction b1 b2 rule: and.induct) |
43141 | 45 |
apply simp_all |
46 |
done |
|
47 |
||
67406 | 48 |
text_raw\<open>\snip{BExpnotdef}{2}{2}{%\<close> |
43141 | 49 |
fun not :: "bexp \<Rightarrow> bexp" where |
45200 | 50 |
"not (Bc True) = Bc False" | |
51 |
"not (Bc False) = Bc True" | |
|
43141 | 52 |
"not b = Not b" |
67406 | 53 |
text_raw\<open>}%endsnip\<close> |
43141 | 54 |
|
45255 | 55 |
lemma bval_not[simp]: "bval (not b) s = (\<not> bval b s)" |
45015 | 56 |
apply(induction b rule: not.induct) |
43141 | 57 |
apply simp_all |
58 |
done |
|
59 |
||
67406 | 60 |
text\<open>Now the overall optimizer:\<close> |
43141 | 61 |
|
67406 | 62 |
text_raw\<open>\snip{BExpbsimpdef}{0}{2}{%\<close> |
43141 | 63 |
fun bsimp :: "bexp \<Rightarrow> bexp" where |
45256 | 64 |
"bsimp (Bc v) = Bc v" | |
65 |
"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
|
66 |
"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
|
67 |
"bsimp (Less a\<^sub>1 a\<^sub>2) = less (asimp a\<^sub>1) (asimp a\<^sub>2)" |
67406 | 68 |
text_raw\<open>}%endsnip\<close> |
43141 | 69 |
|
70 |
value "bsimp (And (Less (N 0) (N 1)) b)" |
|
71 |
||
49920 | 72 |
value "bsimp (And (Less (N 1) (N 0)) (Bc True))" |
43141 | 73 |
|
74 |
theorem "bval (bsimp b) s = bval b s" |
|
45015 | 75 |
apply(induction b) |
43141 | 76 |
apply simp_all |
77 |
done |
|
78 |
||
79 |
end |