theory BExp imports AExp begin subsection "Boolean Expressions" datatype bexp = Bc bool | Not bexp | And bexp bexp | Less aexp aexp text_raw{*\snip{BExpbvaldef}{1}{2}{% *} fun bval :: "bexp ⇒ state ⇒ bool" where "bval (Bc v) s = v" | "bval (Not b) s = (¬ bval b s)" | "bval (And b⇩1 b⇩2) s = (bval b⇩1 s ∧ bval b⇩2 s)" | "bval (Less a⇩1 a⇩2) s = (aval a⇩1 s < aval a⇩2 s)" text_raw{*}%endsnip*} value "bval (Less (V ''x'') (Plus (N 3) (V ''y''))) <''x'' := 3, ''y'' := 1>" subsection "Constant Folding" text{* Optimizing constructors: *} text_raw{*\snip{BExplessdef}{0}{2}{% *} fun less :: "aexp ⇒ aexp ⇒ bexp" where "less (N n⇩1) (N n⇩2) = Bc(n⇩1 < n⇩2)" | "less a⇩1 a⇩2 = Less a⇩1 a⇩2" text_raw{*}%endsnip*} lemma [simp]: "bval (less a1 a2) s = (aval a1 s < aval a2 s)" apply(induction a1 a2 rule: less.induct) apply simp_all done text_raw{*\snip{BExpanddef}{2}{2}{% *} fun "and" :: "bexp ⇒ bexp ⇒ bexp" where "and (Bc True) b = b" | "and b (Bc True) = b" | "and (Bc False) b = Bc False" | "and b (Bc False) = Bc False" | "and b⇩1 b⇩2 = And b⇩1 b⇩2" text_raw{*}%endsnip*} lemma bval_and[simp]: "bval (and b1 b2) s = (bval b1 s ∧ bval b2 s)" apply(induction b1 b2 rule: and.induct) apply simp_all done text_raw{*\snip{BExpnotdef}{2}{2}{% *} fun not :: "bexp ⇒ bexp" where "not (Bc True) = Bc False" | "not (Bc False) = Bc True" | "not b = Not b" text_raw{*}%endsnip*} lemma bval_not[simp]: "bval (not b) s = (¬ bval b s)" apply(induction b rule: not.induct) apply simp_all done text{* Now the overall optimizer: *} text_raw{*\snip{BExpbsimpdef}{0}{2}{% *} fun bsimp :: "bexp ⇒ bexp" where "bsimp (Bc v) = Bc v" | "bsimp (Not b) = not(bsimp b)" | "bsimp (And b⇩1 b⇩2) = and (bsimp b⇩1) (bsimp b⇩2)" | "bsimp (Less a⇩1 a⇩2) = less (asimp a⇩1) (asimp a⇩2)" text_raw{*}%endsnip*} value "bsimp (And (Less (N 0) (N 1)) b)" value "bsimp (And (Less (N 1) (N 0)) (Bc True))" theorem "bval (bsimp b) s = bval b s" apply(induction b) apply simp_all done end