Theory Expressions
section ‹Expressions›
theory Expressions
imports Contracts StateMonad
begin
subsection ‹Semantics of Expressions›
definition lift ::
"(E ⇒ Environment ⇒ CalldataT ⇒ State ⇒ (Stackvalue * Type, Ex, Gas) state_monad)
⇒ (Types ⇒ Types ⇒ Valuetype ⇒ Valuetype ⇒ (Valuetype * Types) option)
⇒ E ⇒ E ⇒ Environment ⇒ CalldataT ⇒ State ⇒ (Stackvalue * Type, Ex, Gas) state_monad"
where
"lift expr f e1 e2 e cd st ≡
(do {
kv1 ← expr e1 e cd st;
(v1, t1) ← case kv1 of (KValue v1, Value t1) ⇒ return (v1, t1) | _ ⇒ (throw Err::(Valuetype * Types, Ex, Gas) state_monad);
kv2 ← expr e2 e cd st;
(v2, t2) ← case kv2 of (KValue v2, Value t2) ⇒ return (v2, t2) | _ ⇒ (throw Err::(Valuetype * Types, Ex, Gas) state_monad);
(v, t) ← (option Err (λ_::Gas. f t1 t2 v1 v2))::(Valuetype * Types, Ex, Gas) state_monad;
return (KValue v, Value t)::(Stackvalue * Type, Ex, Gas) state_monad
})"
declare lift_def[simp, solidity_symbex]
lemma lift_cong [fundef_cong]:
assumes "expr e1 e cd st g = expr' e1 e cd st g"
and "⋀v g'. expr' e1 e cd st g = Normal (v,g') ⟹ expr e2 e cd st g' = expr' e2 e cd st g'"
shows "lift expr f e1 e2 e cd st g = lift expr' f e1 e2 e cd st g"
unfolding lift_def using assms by (auto split: prod.split_asm result.split option.split_asm option.split Stackvalue.split Type.split)