diff --git a/thys/Solidity/Accounts.thy b/thys/Solidity/Accounts.thy --- a/thys/Solidity/Accounts.thy +++ b/thys/Solidity/Accounts.thy @@ -1,267 +1,313 @@ -section\Accounts\ -theory Accounts -imports Valuetypes -begin - -type_synonym Balance = Valuetype - -type_synonym Accounts = "Address \ Balance" - -fun emptyAccount :: "Accounts" -where - "emptyAccount _ = ShowL\<^sub>i\<^sub>n\<^sub>t 0" - -fun accessBalance :: "Accounts \ Address \ Balance" -where - "accessBalance acc ad = acc ad" - -fun updateBalance :: "Address \ Balance \ Accounts \ Accounts" -where - "updateBalance ad bal acc = (\i. (if i = ad then bal else acc i))" - -fun addBalance :: "Address \ Valuetype \ Accounts \ Accounts option" -where - "addBalance ad val acc = - (if ReadL\<^sub>i\<^sub>n\<^sub>t val \ 0 - then (let v = ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance acc ad) + ReadL\<^sub>i\<^sub>n\<^sub>t val - in if (v < 2^256) - then Some (updateBalance ad (ShowL\<^sub>i\<^sub>n\<^sub>t v) acc) - else None) - else None)" - -lemma addBalance_val: - assumes "addBalance ad val acc = Some acc'" - shows "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance acc ad) + ReadL\<^sub>i\<^sub>n\<^sub>t val < 2^256" -using assms by (simp add:Let_def split:if_split_asm) - -lemma addBalance_limit: - assumes "addBalance ad val acc = Some acc'" - and "\ad. ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance acc ad) \ 0 \ ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance acc ad) < 2 ^ 256" - shows "\ad. ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance acc' ad) \ 0 \ ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance acc' ad) < 2 ^ 256" -proof - fix ad' - show "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance acc' ad') \ 0 \ ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance acc' ad') < 2 ^ 256" - proof (cases "ReadL\<^sub>i\<^sub>n\<^sub>t val \ 0") - case t1: True - define v where v_def: "v = ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance acc ad) + ReadL\<^sub>i\<^sub>n\<^sub>t val" - with assms(2) have "v \0" by (simp add: t1) - then show ?thesis - proof (cases "v < 2^256") - case t2: True - with t1 v_def have "addBalance ad val acc = Some (updateBalance ad (ShowL\<^sub>i\<^sub>n\<^sub>t v) acc)" by simp - with t2 `v \0` show ?thesis using assms Read_ShowL_id by auto - next - case False - with t1 v_def show ?thesis using assms(1) by simp - qed - next - case False - then show ?thesis using assms(1) by simp - qed -qed - -lemma addBalance_add: - assumes "addBalance ad val acc = Some acc'" - shows "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance acc' ad) = ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance acc ad) + ReadL\<^sub>i\<^sub>n\<^sub>t val" -proof - - define v where "v = ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance acc ad) + ReadL\<^sub>i\<^sub>n\<^sub>t val" - with assms have "acc' = updateBalance ad (ShowL\<^sub>i\<^sub>n\<^sub>t v) acc" by (simp add:Let_def split:if_split_asm) - thus ?thesis using v_def Read_ShowL_id assms by (simp split:if_split_asm) -qed - -lemma addBalance_mono: - assumes "addBalance ad val acc = Some acc'" - shows "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance acc' ad) \ ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance acc ad)" -proof - - define v where "v = ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance acc ad) + ReadL\<^sub>i\<^sub>n\<^sub>t val" - with assms have "acc' = updateBalance ad (ShowL\<^sub>i\<^sub>n\<^sub>t v) acc" by (simp add:Let_def split:if_split_asm) - thus ?thesis using v_def Read_ShowL_id assms by (simp split:if_split_asm) -qed - -lemma addBalance_eq: - assumes "addBalance ad val acc = Some acc'" - and "ad \ ad'" - shows "(accessBalance acc ad') = (accessBalance acc' ad')" -proof - - define v where "v = ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance acc ad) + ReadL\<^sub>i\<^sub>n\<^sub>t val" - with assms have "acc' = updateBalance ad (ShowL\<^sub>i\<^sub>n\<^sub>t v) acc" by (simp add:Let_def split:if_split_asm) - thus ?thesis using v_def Read_ShowL_id assms by (simp split:if_split_asm) -qed - -fun subBalance :: "Address \ Valuetype \ Accounts \ Accounts option" - where - "subBalance ad val acc = - (if ReadL\<^sub>i\<^sub>n\<^sub>t val \ 0 - then (let v = ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance acc ad) - ReadL\<^sub>i\<^sub>n\<^sub>t val - in if (v \ 0) - then Some (updateBalance ad (ShowL\<^sub>i\<^sub>n\<^sub>t v) acc) - else None) - else None)" - -lemma subBalance_val: - assumes "subBalance ad val acc = Some acc'" - shows "ReadL\<^sub>i\<^sub>n\<^sub>t val \ 0" -using assms by (simp split:if_split_asm) - -lemma subBalance_sub: - assumes "subBalance ad val acc = Some acc'" - shows "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance acc' ad) = ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance acc ad) - ReadL\<^sub>i\<^sub>n\<^sub>t val" -proof - - define v where "v = ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance acc ad) - ReadL\<^sub>i\<^sub>n\<^sub>t val" - with assms have "acc' = updateBalance ad (ShowL\<^sub>i\<^sub>n\<^sub>t v) acc" by (simp add:Let_def split:if_split_asm) - thus ?thesis using v_def Read_ShowL_id assms by (simp split:if_split_asm) -qed - -lemma subBalance_limit: - assumes "subBalance ad val acc = Some acc'" - and "\ad. ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance acc ad) \ 0 \ ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance acc ad) < 2 ^ 256" - shows "\ad. ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance acc' ad) \ 0 \ ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance acc' ad) < 2 ^ 256" -proof - fix ad' - show "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance acc' ad') \ 0 \ ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance acc' ad') < 2 ^ 256" - proof (cases "ReadL\<^sub>i\<^sub>n\<^sub>t val \ 0") - case t1: True - define v where v_def: "v = ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance acc ad) - ReadL\<^sub>i\<^sub>n\<^sub>t val" - with assms(2) t1 have "v < 2 ^ 256" by (smt (verit)) - then show ?thesis - proof (cases "v \ 0") - case t2: True - with t1 v_def have "subBalance ad val acc = Some (updateBalance ad (ShowL\<^sub>i\<^sub>n\<^sub>t v) acc)" by simp - with t2 `v < 2 ^ 256` show ?thesis using assms Read_ShowL_id by auto - next - case False - with t1 v_def show ?thesis using assms(1) by simp - qed - next - case False - then show ?thesis using assms(1) by simp - qed -qed - -lemma subBalance_mono: - assumes "subBalance ad val acc = Some acc'" - shows "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance acc ad) \ ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance acc' ad)" -proof - - define v where "v = ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance acc ad) - ReadL\<^sub>i\<^sub>n\<^sub>t val" - with assms have "acc' = updateBalance ad (ShowL\<^sub>i\<^sub>n\<^sub>t v) acc" by (simp add:Let_def split:if_split_asm) - thus ?thesis using v_def Read_ShowL_id assms by (simp split:if_split_asm) -qed - -lemma subBalance_eq: - assumes "subBalance ad val acc = Some acc'" - and "ad \ ad'" - shows "(accessBalance acc ad') = (accessBalance acc' ad')" -proof - - define v where "v = ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance acc ad) - ReadL\<^sub>i\<^sub>n\<^sub>t val" - with assms have "acc' = updateBalance ad (ShowL\<^sub>i\<^sub>n\<^sub>t v) acc" by (simp add:Let_def split:if_split_asm) - thus ?thesis using v_def Read_ShowL_id assms by (simp split:if_split_asm) -qed - -fun transfer :: "Address \ Address \ Valuetype \ Accounts \ Accounts option" -where - "transfer ads addr val acc = - (case subBalance ads val acc of - Some acc' \ addBalance addr val acc' - | None \ None)" - -lemma transfer_val1: - assumes "transfer ads addr val acc = Some acc'" - shows "ReadL\<^sub>i\<^sub>n\<^sub>t val \ 0" -proof - - from assms(1) obtain acc'' - where *: "subBalance ads val acc = Some acc''" - and **: "addBalance addr val acc'' = Some acc'" by (simp split:if_split_asm) - then show "ReadL\<^sub>i\<^sub>n\<^sub>t val \ 0" using subBalance_val[OF *] by simp -qed - -lemma transfer_val2: - assumes "transfer ads addr val acc = Some acc'" - and "ads \ addr" - shows "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance acc addr) + ReadL\<^sub>i\<^sub>n\<^sub>t val < 2^256" -proof - - from assms(1) obtain acc'' - where *: "subBalance ads val acc = Some acc''" - and **: "addBalance addr val acc'' = Some acc'" by (simp split:if_split_asm) - - then have "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance acc'' addr) + ReadL\<^sub>i\<^sub>n\<^sub>t val < 2^256" using addBalance_val[OF **] by simp - with * show ?thesis using assms(2) subBalance_eq[OF *] by simp -qed - -lemma transfer_add: - assumes "transfer ads addr val acc = Some acc'" - and "addr \ ads" - shows "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance acc' addr) = ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance acc addr) + ReadL\<^sub>i\<^sub>n\<^sub>t val" -proof - - from assms(1) obtain acc'' - where *: "subBalance ads val acc = Some acc''" - and **: "addBalance addr val acc'' = Some acc'" by (simp split:if_split_asm) - - with assms(2) have "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance acc addr) = ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance acc'' addr)" using subBalance_eq[OF *] by simp - moreover have "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance acc' addr) = ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance acc'' addr) + ReadL\<^sub>i\<^sub>n\<^sub>t val" using addBalance_add[OF **] by simp - ultimately show ?thesis using Read_ShowL_id by simp -qed - -lemma transfer_sub: - assumes "transfer ads addr val acc = Some acc'" - and "addr \ ads" - shows "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance acc' ads) = ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance acc ads) - ReadL\<^sub>i\<^sub>n\<^sub>t val" -proof - - from assms(1) obtain acc'' - where *: "subBalance ads val acc = Some acc''" - and **: "addBalance addr val acc'' = Some acc'" by (simp split:if_split_asm) - - then have "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance acc'' ads) = ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance acc ads) - ReadL\<^sub>i\<^sub>n\<^sub>t val" using subBalance_sub[OF *] by simp - moreover from assms(2) have "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance acc' ads) = ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance acc'' ads)" using addBalance_eq[OF **] by simp - ultimately show ?thesis using Read_ShowL_id by simp -qed - -lemma transfer_mono: - assumes "transfer ads addr val acc = Some acc'" - shows "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance acc' addr) \ ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance acc addr)" -proof - - from assms(1) obtain acc'' - where *: "subBalance ads val acc = Some acc''" - and **: "addBalance addr val acc'' = Some acc'" by (simp split:if_split_asm) - - show ?thesis - proof (cases "addr = ads") - case True - with * have "acc'' = updateBalance addr (ShowL\<^sub>i\<^sub>n\<^sub>t (ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance acc addr) - ReadL\<^sub>i\<^sub>n\<^sub>t val)) acc" by (simp add:Let_def split: if_split_asm) - moreover from ** have "acc'=updateBalance addr (ShowL\<^sub>i\<^sub>n\<^sub>t (ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance acc'' addr) + ReadL\<^sub>i\<^sub>n\<^sub>t val)) acc''" by (simp add:Let_def split: if_split_asm) - ultimately show ?thesis using Read_ShowL_id by auto - next - case False - then have "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance acc addr) \ ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance acc'' addr)" using subBalance_eq[OF *] by simp - also have "\ \ ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance acc' addr)" using addBalance_mono[OF **] by simp - finally show ?thesis . - qed -qed - -lemma transfer_eq: - assumes "transfer ads addr val acc = Some acc'" - and "ad \ ads" - and "ad \ addr" - shows "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance acc' ad) = ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance acc ad)" -proof - - from assms(1) obtain acc'' - where *: "subBalance ads val acc = Some acc''" - and **: "addBalance addr val acc'' = Some acc'" by (simp split:if_split_asm) - - from assms(3) have "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance acc' ad) = ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance acc'' ad)" using addBalance_eq[OF **] by simp - also from assms(2) have "\ = ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance acc ad)" using subBalance_eq[OF *] by simp - finally show ?thesis . -qed - -lemma transfer_limit: - assumes "transfer ads addr val acc = Some acc'" - and "\ad. ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance acc ad) \ 0 \ ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance acc ad) < 2 ^ 256" - shows "\ad. ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance acc' ad) \ 0 \ ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance acc' ad) < 2 ^ 256" -proof - fix ad' - from assms(1) obtain acc'' where "subBalance ads val acc = Some acc''" and "addBalance addr val acc'' = Some acc'" by (simp split: if_split_asm) - with subBalance_limit[OF _ assms(2)] - show "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance acc' ad') \ 0 \ ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance acc' ad') < 2 ^ 256" - using addBalance_limit by presburger -qed - +section\Accounts\ +theory Accounts +imports Valuetypes +begin + +type_synonym Balance = Valuetype +type_synonym Identifier = String.literal + +(* +Contract None means not initialized yet +*) +datatype atype = + EOA + | Contract Identifier + +record account = + bal :: Balance + type :: "atype option" + contracts :: nat + +lemma bind_case_atype_cong [fundef_cong]: + assumes "x = x'" + and "x = EOA \ f s = f' s" + and "\a. x = Contract a \ g a s = g' a s" + shows "(case x of EOA \ f | Contract a \ g a) s + = (case x' of EOA \ f' | Contract a \ g' a) s" + using assms by (cases x, auto) + +definition emptyAcc :: account + where "emptyAcc = \bal = ShowL\<^sub>i\<^sub>n\<^sub>t 0, type = None, contracts = 0\" + +type_synonym Accounts = "Address \ account" + +definition emptyAccount :: "Accounts" +where + "emptyAccount _ = emptyAcc" + +definition addBalance :: "Address \ Valuetype \ Accounts \ Accounts option" +where + "addBalance ad val acc = + (if ReadL\<^sub>i\<^sub>n\<^sub>t val \ 0 + then (let v = ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc ad)) + ReadL\<^sub>i\<^sub>n\<^sub>t val + in if (v < 2^256) + then Some (acc(ad := acc ad \bal:=ShowL\<^sub>i\<^sub>n\<^sub>t v\)) + else None) + else None)" + +lemma addBalance_val1: + assumes "addBalance ad val acc = Some acc'" + shows "ReadL\<^sub>i\<^sub>n\<^sub>t val \ 0" +using assms unfolding addBalance_def by (simp add:Let_def split:if_split_asm) + +lemma addBalance_val2: + assumes "addBalance ad val acc = Some acc'" + shows "ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc ad)) + ReadL\<^sub>i\<^sub>n\<^sub>t val < 2^256" +using assms unfolding addBalance_def by (simp add:Let_def split:if_split_asm) + +lemma addBalance_limit: + assumes "addBalance ad val acc = Some acc'" + and "\ad. ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc ad)) \ 0 \ ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc ad)) < 2 ^ 256" + shows "\ad. ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc' ad)) \ 0 \ ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc' ad)) < 2 ^ 256" +proof + fix ad' + show "ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc' ad')) \ 0 \ ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc' ad')) < 2 ^ 256" + proof (cases "ReadL\<^sub>i\<^sub>n\<^sub>t val \ 0") + case t1: True + define v where v_def: "v = ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc ad)) + ReadL\<^sub>i\<^sub>n\<^sub>t val" + with assms(2) have "v \0" by (simp add: t1) + then show ?thesis + proof (cases "v < 2^256") + case t2: True + with t1 v_def have "addBalance ad val acc = Some (acc(ad := acc ad\bal:=ShowL\<^sub>i\<^sub>n\<^sub>t v\))" unfolding addBalance_def by simp + with t2 `v \0` show ?thesis using assms Read_ShowL_id by auto + next + case False + with t1 v_def show ?thesis using assms(1) unfolding addBalance_def by simp + qed + next + case False + then show ?thesis using assms(1) unfolding addBalance_def by simp + qed +qed + +lemma addBalance_add: + assumes "addBalance ad val acc = Some acc'" + shows "ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc' ad)) = ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc ad)) + ReadL\<^sub>i\<^sub>n\<^sub>t val" +proof - + define v where "v = ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc ad)) + ReadL\<^sub>i\<^sub>n\<^sub>t val" + with assms have "acc' = acc(ad := acc ad\bal:=ShowL\<^sub>i\<^sub>n\<^sub>t v\)" unfolding addBalance_def by (simp add:Let_def split:if_split_asm) + thus ?thesis using v_def Read_ShowL_id assms by (simp split:if_split_asm) +qed + +lemma addBalance_mono: + assumes "addBalance ad val acc = Some acc'" + shows "ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc' ad)) \ ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc ad))" +proof - + define v where "v = ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc ad)) + ReadL\<^sub>i\<^sub>n\<^sub>t val" + with assms have "acc' = acc(ad := acc ad\bal:=ShowL\<^sub>i\<^sub>n\<^sub>t v\)" unfolding addBalance_def by (simp add:Let_def split:if_split_asm) + thus ?thesis using v_def Read_ShowL_id assms unfolding addBalance_def by (simp split:if_split_asm) +qed + +lemma addBalance_eq: + assumes "addBalance ad val acc = Some acc'" + and "ad \ ad'" + shows "bal (acc ad') = bal (acc' ad')" +proof - + define v where "v = ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc ad)) + ReadL\<^sub>i\<^sub>n\<^sub>t val" + with assms have "acc' = acc(ad := acc ad\bal:=ShowL\<^sub>i\<^sub>n\<^sub>t v\)" unfolding addBalance_def by (simp add:Let_def split:if_split_asm) + thus ?thesis using v_def Read_ShowL_id assms by (simp split:if_split_asm) +qed + +definition subBalance :: "Address \ Valuetype \ Accounts \ Accounts option" + where + "subBalance ad val acc = + (if ReadL\<^sub>i\<^sub>n\<^sub>t val \ 0 + then (let v = ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc ad)) - ReadL\<^sub>i\<^sub>n\<^sub>t val + in if (v \ 0) + then Some (acc(ad := acc ad\bal:=ShowL\<^sub>i\<^sub>n\<^sub>t v\)) + else None) + else None)" + +lemma subBalance_val1: + assumes "subBalance ad val acc = Some acc'" + shows "ReadL\<^sub>i\<^sub>n\<^sub>t val \ 0" +using assms unfolding subBalance_def by (simp split:if_split_asm) + +lemma subBalance_val2: + assumes "subBalance ad val acc = Some acc'" + shows "ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc ad)) - ReadL\<^sub>i\<^sub>n\<^sub>t val \ 0" +using assms unfolding subBalance_def by (simp split:if_split_asm) + +lemma subBalance_sub: + assumes "subBalance ad val acc = Some acc'" + shows "ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc' ad)) = ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc ad)) - ReadL\<^sub>i\<^sub>n\<^sub>t val" +proof - + define v where "v = ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc ad)) - ReadL\<^sub>i\<^sub>n\<^sub>t val" + with assms have "acc' = acc(ad := acc ad\bal:=ShowL\<^sub>i\<^sub>n\<^sub>t v\)" unfolding subBalance_def by (simp add:Let_def split:if_split_asm) + thus ?thesis using v_def Read_ShowL_id assms by (simp split:if_split_asm) +qed + +lemma subBalance_limit: + assumes "subBalance ad val acc = Some acc'" + and "\ad. ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc ad)) \ 0 \ ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc ad)) < 2 ^ 256" + shows "\ad. ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc' ad)) \ 0 \ ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc' ad)) < 2 ^ 256" +proof + fix ad' + show "ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc' ad')) \ 0 \ ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc' ad')) < 2 ^ 256" + proof (cases "ReadL\<^sub>i\<^sub>n\<^sub>t val \ 0") + case t1: True + define v where v_def: "v = ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc ad)) - ReadL\<^sub>i\<^sub>n\<^sub>t val" + with assms(2) t1 have "v < 2 ^ 256" by (smt (verit)) + then show ?thesis + proof (cases "v \ 0") + case t2: True + with t1 v_def have "subBalance ad val acc = Some (acc(ad := acc ad\bal:=ShowL\<^sub>i\<^sub>n\<^sub>t v\))" unfolding subBalance_def by simp + with t2 `v < 2 ^ 256` show ?thesis using assms Read_ShowL_id by auto + next + case False + with t1 v_def show ?thesis using assms(1) unfolding subBalance_def by simp + qed + next + case False + then show ?thesis using assms(1) unfolding subBalance_def by simp + qed +qed + +lemma subBalance_mono: + assumes "subBalance ad val acc = Some acc'" + shows "ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc ad)) \ ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc' ad))" +proof - + define v where "v = ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc ad)) - ReadL\<^sub>i\<^sub>n\<^sub>t val" + with assms have "acc' = acc(ad := acc ad\bal:=ShowL\<^sub>i\<^sub>n\<^sub>t v\)" unfolding subBalance_def by (simp add:Let_def split:if_split_asm) + thus ?thesis using v_def Read_ShowL_id assms unfolding subBalance_def by (simp split:if_split_asm) +qed + +lemma subBalance_eq: + assumes "subBalance ad val acc = Some acc'" + and "ad \ ad'" + shows "(bal (acc ad')) = (bal (acc' ad'))" +proof - + define v where "v = ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc ad)) - ReadL\<^sub>i\<^sub>n\<^sub>t val" + with assms have "acc' = acc(ad := acc ad\bal:=ShowL\<^sub>i\<^sub>n\<^sub>t v\)" unfolding subBalance_def by (simp add:Let_def split:if_split_asm) + thus ?thesis using v_def Read_ShowL_id assms by (simp split:if_split_asm) +qed + +definition transfer :: "Address \ Address \ Valuetype \ Accounts \ Accounts option" +where + "transfer ads addr val acc = + (case subBalance ads val acc of + Some acc' \ addBalance addr val acc' + | None \ None)" + +lemma transfer_val1: + assumes "transfer ads addr val acc = Some acc'" + shows "ReadL\<^sub>i\<^sub>n\<^sub>t val \ 0" +proof - + from assms(1) obtain acc'' + where *: "subBalance ads val acc = Some acc''" + and **: "addBalance addr val acc'' = Some acc'" by (simp add: subBalance_def transfer_def split:if_split_asm) + then show "ReadL\<^sub>i\<^sub>n\<^sub>t val \ 0" using subBalance_val1[OF *] by simp +qed + +lemma transfer_val2: + assumes "transfer ads addr val acc = Some acc'" + and "ads \ addr" + shows "ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc addr)) + ReadL\<^sub>i\<^sub>n\<^sub>t val < 2^256" +proof - + from assms(1) obtain acc'' + where *: "subBalance ads val acc = Some acc''" + and **: "addBalance addr val acc'' = Some acc'" by (simp add: subBalance_def transfer_def split:if_split_asm) + + have "ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc'' addr)) + ReadL\<^sub>i\<^sub>n\<^sub>t val < 2^256" using addBalance_val2[OF **] by simp + with * show ?thesis using assms(2) subBalance_eq[OF *] by simp +qed + +lemma transfer_val3: + assumes "transfer ads addr val acc = Some acc'" + shows "ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc ads)) - ReadL\<^sub>i\<^sub>n\<^sub>t val \ 0" +using assms by (auto simp add: Let_def subBalance_def transfer_def split:if_split_asm) + +lemma transfer_add: + assumes "transfer ads addr val acc = Some acc'" + and "addr \ ads" + shows "ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc' addr)) = ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc addr)) + ReadL\<^sub>i\<^sub>n\<^sub>t val" +proof - + from assms(1) obtain acc'' + where *: "subBalance ads val acc = Some acc''" + and **: "addBalance addr val acc'' = Some acc'" by (simp add: subBalance_def transfer_def split:if_split_asm) + + with assms(2) have "ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc addr)) = ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc'' addr))" using subBalance_eq[OF *] by simp + moreover have "ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc' addr)) = ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc'' addr)) + ReadL\<^sub>i\<^sub>n\<^sub>t val" using addBalance_add[OF **] by simp + ultimately show ?thesis using Read_ShowL_id by simp +qed + +lemma transfer_sub: + assumes "transfer ads addr val acc = Some acc'" + and "addr \ ads" + shows "ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc' ads)) = ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc ads)) - ReadL\<^sub>i\<^sub>n\<^sub>t val" +proof - + from assms(1) obtain acc'' + where *: "subBalance ads val acc = Some acc''" + and **: "addBalance addr val acc'' = Some acc'" by (simp add: subBalance_def transfer_def split:if_split_asm) + + then have "ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc'' ads)) = ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc ads)) - ReadL\<^sub>i\<^sub>n\<^sub>t val" using subBalance_sub[OF *] by simp + moreover from assms(2) have "ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc' ads)) = ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc'' ads))" using addBalance_eq[OF **] by simp + ultimately show ?thesis using Read_ShowL_id by simp +qed + +lemma transfer_same: + assumes "transfer ad ad' val acc = Some acc'" + and "ad = ad'" + shows "ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc ad)) = ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc' ad))" +proof - + from assms obtain acc'' where *: "subBalance ad val acc = Some acc''" by (simp add: subBalance_def transfer_def split:if_split_asm) + with assms have **: "addBalance ad val acc'' = Some acc'" by (simp add: transfer_def split:if_split_asm) + moreover from * have "ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc'' ad)) = ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc ad)) - ReadL\<^sub>i\<^sub>n\<^sub>t val" using subBalance_sub by simp + moreover from ** have "ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc' ad)) = ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc'' ad)) + ReadL\<^sub>i\<^sub>n\<^sub>t val" using addBalance_add by simp + ultimately show ?thesis by simp +qed + +lemma transfer_mono: + assumes "transfer ads addr val acc = Some acc'" + shows "ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc' addr)) \ ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc addr))" +proof - + from assms(1) obtain acc'' + where *: "subBalance ads val acc = Some acc''" + and **: "addBalance addr val acc'' = Some acc'" by (simp add: subBalance_def transfer_def split:if_split_asm) + + show ?thesis + proof (cases "addr = ads") + case True + with * have "acc'' = acc(addr:=acc addr\bal := ShowL\<^sub>i\<^sub>n\<^sub>t (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc addr)) - ReadL\<^sub>i\<^sub>n\<^sub>t val)\)" by (simp add:Let_def subBalance_def split: if_split_asm) + moreover from ** have "acc'=acc''(addr:=acc'' addr\bal := ShowL\<^sub>i\<^sub>n\<^sub>t (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc'' addr)) + ReadL\<^sub>i\<^sub>n\<^sub>t val)\)" unfolding addBalance_def by (simp add:Let_def split: if_split_asm) + ultimately show ?thesis using Read_ShowL_id by auto + next + case False + then have "ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc addr)) \ ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc'' addr))" using subBalance_eq[OF *] by simp + also have "\ \ ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc' addr))" using addBalance_mono[OF **] by simp + finally show ?thesis . + qed +qed + +lemma transfer_eq: + assumes "transfer ads addr val acc = Some acc'" + and "ad \ ads" + and "ad \ addr" + shows "bal (acc' ad) = bal (acc ad)" +using assms by (auto simp add: Let_def addBalance_def subBalance_def transfer_def split:if_split_asm) + +lemma transfer_limit: + assumes "transfer ads addr val acc = Some acc'" + and "\ad. ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc ad)) \ 0 \ ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc ad)) < 2 ^ 256" + shows "\ad. ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc' ad)) \ 0 \ ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc' ad)) < 2 ^ 256" +proof + fix ad' + from assms(1) obtain acc'' where "subBalance ads val acc = Some acc''" and "addBalance addr val acc'' = Some acc'" by (simp add: subBalance_def transfer_def split: if_split_asm) + with subBalance_limit[OF _ assms(2)] + show "ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc' ad')) \ 0 \ ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc' ad')) < 2 ^ 256" + using addBalance_limit by presburger +qed + +lemma transfer_type_same: + assumes "transfer ads addr val acc = Some acc'" + shows "type (acc' ad) = type (acc ad)" +using assms by (auto simp add: Let_def addBalance_def subBalance_def transfer_def split:if_split_asm) + +lemma transfer_contracts_same: + assumes "transfer ads addr val acc = Some acc'" + shows "contracts (acc' ad) = contracts (acc ad)" +using assms by (auto simp add: Let_def addBalance_def subBalance_def transfer_def split:if_split_asm) + + end \ No newline at end of file diff --git a/thys/Solidity/Compile_Evaluator.thy b/thys/Solidity/Compile_Evaluator.thy --- a/thys/Solidity/Compile_Evaluator.thy +++ b/thys/Solidity/Compile_Evaluator.thy @@ -1,41 +1,40 @@ -section\Generating an Exectuable of the Evaluator\ - -theory - Compile_Evaluator - imports - Solidity_Evaluator -begin - - -compile_generated_files _ (in Solidity_Evaluator) export_files "solidity-evaluator" (executable) - where \fn dir => - let - val modules_src = - Generated_Files.get_files \<^theory>\Solidity_Evaluator\ - |> filter (fn p => String.isSubstring "src" (Path.implode (#path p))) - |> map (#path #> Path.implode #> unsuffix ".hs" #> space_explode "/" #> space_implode "." - #> unprefix "code.solidity-evaluator.src."); - val modules_app = - Generated_Files.get_files \<^theory>\Solidity_Evaluator\ - |> filter (fn p => String.isSubstring "app" (Path.implode (#path p))) - |> map (#path #> Path.implode #> unsuffix ".hs" #> space_explode "/" #> space_implode "." - #> unprefix "code.solidity-evaluator.app."); - val _ = - GHC.new_project dir - {name = "solidity-evaluator", - depends = - [], - modules = modules_app}; - val _ = writeln (Path.implode dir) - val res = Generated_Files.execute dir \Build\ (String.concat [ - "echo \"\n default-extensions: TypeSynonymInstances, FlexibleInstances\" >> solidity-evaluator.cabal" - ," && rm -rf src" - ," && mv code/solidity-evaluator/src src" - ," && mv code/solidity-evaluator/app/* src/" - ," && isabelle ghc_stack install --local-bin-path . `pwd`"]) - in - writeln (res) - end\ - - -end +section\Generating an Exectuable of the Evaluator\ + +theory + Compile_Evaluator + imports + Solidity_Evaluator +begin + + +compile_generated_files _ (in Solidity_Evaluator) export_files "solidity-evaluator" (exe) + where \fn dir => + let + val modules_src = + Generated_Files.get_files \<^theory>\Solidity_Evaluator\ + |> filter (fn p => String.isSubstring "src" (Path.implode (#path p))) + |> map (#path #> Path.implode #> unsuffix ".hs" #> space_explode "/" #> space_implode "." + #> unprefix "code.solidity-evaluator.src."); + val modules_app = + Generated_Files.get_files \<^theory>\Solidity_Evaluator\ + |> filter (fn p => String.isSubstring "app" (Path.implode (#path p))) + |> map (#path #> Path.implode #> unsuffix ".hs" #> space_explode "/" #> space_implode "." + #> unprefix "code.solidity-evaluator.app."); + val _ = + GHC.new_project dir + {name = "solidity-evaluator", + depends = + [], + modules = modules_app}; + val res = Generated_Files.execute dir \Build\ (String.concat [ + "echo \"\n default-extensions: TypeSynonymInstances, FlexibleInstances\" >> solidity-evaluator.cabal" + ," && rm -rf src" + ," && mv code/solidity-evaluator/src src" + ," && mv code/solidity-evaluator/app/* src/" + ," && isabelle ghc_stack install --local-bin-path . ."]) + in + writeln (res) + end\ + + +end diff --git a/thys/Solidity/Constant_Folding.thy b/thys/Solidity/Constant_Folding.thy --- a/thys/Solidity/Constant_Folding.thy +++ b/thys/Solidity/Constant_Folding.thy @@ -1,6016 +1,6203 @@ -section\Constant Folding\ -theory Constant_Folding -imports - Solidity_Main -begin - -text\ - The following function optimizes expressions w.r.t. gas consumption. -\ -primrec eupdate :: "E \ E" -and lupdate :: "L \ L" -where - "lupdate (Id i) = Id i" -| "lupdate (Ref i exp) = Ref i (map eupdate exp)" -| "eupdate (E.INT b v) = - (if (b\vbits) - then if v \ 0 - then E.INT b (-(2^(b-1)) + (v+2^(b-1)) mod (2^b)) - else E.INT b (2^(b-1) - (-v+2^(b-1)-1) mod (2^b) - 1) - else E.INT b v)" -| "eupdate (UINT b v) = (if (b\vbits) then UINT b (v mod (2^b)) else UINT b v)" -| "eupdate (ADDRESS a) = ADDRESS a" -| "eupdate (BALANCE a) = BALANCE a" -| "eupdate THIS = THIS" -| "eupdate SENDER = SENDER" -| "eupdate VALUE = VALUE" -| "eupdate TRUE = TRUE" -| "eupdate FALSE = FALSE" -| "eupdate (LVAL l) = LVAL (lupdate l)" -| "eupdate (PLUS ex1 ex2) = - (case (eupdate ex1) of - E.INT b1 v1 \ - if b1 \ vbits - then (case (eupdate ex2) of - E.INT b2 v2 \ - if b2\vbits - then let v=v1+v2 in - if v \ 0 - then E.INT (max b1 b2) (-(2^((max b1 b2)-1)) + (v+2^((max b1 b2)-1)) mod (2^(max b1 b2))) - else E.INT (max b1 b2) (2^((max b1 b2)-1) - (-v+2^((max b1 b2)-1)-1) mod (2^(max b1 b2)) - 1) - else (PLUS (E.INT b1 v1) (E.INT b2 v2)) - | UINT b2 v2 \ - if b2\vbits \ b2 < b1 - then let v=v1+v2 in - if v \ 0 - then E.INT b1 (-(2^(b1-1)) + (v+2^(b1-1)) mod (2^b1)) - else E.INT b1 (2^(b1-1) - (-v+2^(b1-1)-1) mod (2^b1) - 1) - else PLUS (E.INT b1 v1) (UINT b2 v2) - | _ \ PLUS (E.INT b1 v1) (eupdate ex2)) - else PLUS (E.INT b1 v1) (eupdate ex2) - | UINT b1 v1 \ - if b1 \ vbits - then (case (eupdate ex2) of - UINT b2 v2 \ - if b2 \ vbits - then UINT (max b1 b2) ((v1 + v2) mod (2^(max b1 b2))) - else (PLUS (UINT b1 v1) (UINT b2 v2)) - | E.INT b2 v2 \ - if b2\vbits \ b1 < b2 - then let v=v1+v2 in - if v \ 0 - then E.INT b2 (-(2^(b2-1)) + (v+2^(b2-1)) mod (2^b2)) - else E.INT b2 (2^(b2-1) - (-v+2^(b2-1)-1) mod (2^b2) - 1) - else PLUS (UINT b1 v1) (E.INT b2 v2) - | _ \ PLUS (UINT b1 v1) (eupdate ex2)) - else PLUS (UINT b1 v1) (eupdate ex2) - | _ \ PLUS (eupdate ex1) (eupdate ex2))" -| "eupdate (MINUS ex1 ex2) = - (case (eupdate ex1) of - E.INT b1 v1 \ - if b1 \ vbits - then (case (eupdate ex2) of - E.INT b2 v2 \ - if b2\vbits - then let v=v1-v2 in - if v \ 0 - then E.INT (max b1 b2) (-(2^((max b1 b2)-1)) + (v+2^((max b1 b2)-1)) mod (2^(max b1 b2))) - else E.INT (max b1 b2) (2^((max b1 b2)-1) - (-v+2^((max b1 b2)-1)-1) mod (2^(max b1 b2)) - 1) - else (MINUS (E.INT b1 v1) (E.INT b2 v2)) - | UINT b2 v2 \ - if b2\vbits \ b2 < b1 - then let v=v1-v2 in - if v \ 0 - then E.INT b1 (-(2^(b1-1)) + (v+2^(b1-1)) mod (2^b1)) - else E.INT b1 (2^(b1-1) - (-v+2^(b1-1)-1) mod (2^b1) - 1) - else MINUS (E.INT b1 v1) (UINT b2 v2) - | _ \ MINUS (E.INT b1 v1) (eupdate ex2)) - else MINUS (E.INT b1 v1) (eupdate ex2) - | UINT b1 v1 \ - if b1 \ vbits - then (case (eupdate ex2) of - UINT b2 v2 \ - if b2 \ vbits - then UINT (max b1 b2) ((v1 - v2) mod (2^(max b1 b2))) - else (MINUS (UINT b1 v1) (UINT b2 v2)) - | E.INT b2 v2 \ - if b2\vbits \ b1 < b2 - then let v=v1-v2 in - if v \ 0 - then E.INT b2 (-(2^(b2-1)) + (v+2^(b2-1)) mod (2^b2)) - else E.INT b2 (2^(b2-1) - (-v+2^(b2-1)-1) mod (2^b2) - 1) - else MINUS (UINT b1 v1) (E.INT b2 v2) - | _ \ MINUS (UINT b1 v1) (eupdate ex2)) - else MINUS (UINT b1 v1) (eupdate ex2) - | _ \ MINUS (eupdate ex1) (eupdate ex2))" -| "eupdate (EQUAL ex1 ex2) = - (case (eupdate ex1) of - E.INT b1 v1 \ - if b1 \ vbits - then (case (eupdate ex2) of - E.INT b2 v2 \ - if b2\vbits - then if v1 = v2 - then TRUE - else FALSE - else EQUAL (E.INT b1 v1) (E.INT b2 v2) - | UINT b2 v2 \ - if b2\vbits \ b2 < b1 - then if v1 = v2 - then TRUE - else FALSE - else EQUAL (E.INT b1 v1) (UINT b2 v2) - | _ \ EQUAL (E.INT b1 v1) (eupdate ex2)) - else EQUAL (E.INT b1 v1) (eupdate ex2) - | UINT b1 v1 \ - if b1 \ vbits - then (case (eupdate ex2) of - UINT b2 v2 \ - if b2 \ vbits - then if v1 = v2 - then TRUE - else FALSE - else EQUAL (E.INT b1 v1) (UINT b2 v2) - | E.INT b2 v2 \ - if b2\vbits \ b1 < b2 - then if v1 = v2 - then TRUE - else FALSE - else EQUAL (UINT b1 v1) (E.INT b2 v2) - | _ \ EQUAL (UINT b1 v1) (eupdate ex2)) - else EQUAL (UINT b1 v1) (eupdate ex2) - | _ \ EQUAL (eupdate ex1) (eupdate ex2))" -| "eupdate (LESS ex1 ex2) = - (case (eupdate ex1) of - E.INT b1 v1 \ - if b1 \ vbits - then (case (eupdate ex2) of - E.INT b2 v2 \ - if b2\vbits - then if v1 < v2 - then TRUE - else FALSE - else LESS (E.INT b1 v1) (E.INT b2 v2) - | UINT b2 v2 \ - if b2\vbits \ b2 < b1 - then if v1 < v2 - then TRUE - else FALSE - else LESS (E.INT b1 v1) (UINT b2 v2) - | _ \ LESS (E.INT b1 v1) (eupdate ex2)) - else LESS (E.INT b1 v1) (eupdate ex2) - | UINT b1 v1 \ - if b1 \ vbits - then (case (eupdate ex2) of - UINT b2 v2 \ - if b2 \ vbits - then if v1 < v2 - then TRUE - else FALSE - else LESS (E.INT b1 v1) (UINT b2 v2) - | E.INT b2 v2 \ - if b2\vbits \ b1 < b2 - then if v1 < v2 - then TRUE - else FALSE - else LESS (UINT b1 v1) (E.INT b2 v2) - | _ \ LESS (UINT b1 v1) (eupdate ex2)) - else LESS (UINT b1 v1) (eupdate ex2) - | _ \ LESS (eupdate ex1) (eupdate ex2))" -| "eupdate (AND ex1 ex2) = - (case (eupdate ex1) of - TRUE \ (case (eupdate ex2) of - TRUE \ TRUE - | FALSE \ FALSE - | _ \ AND TRUE (eupdate ex2)) - | FALSE \ (case (eupdate ex2) of - TRUE \ FALSE - | FALSE \ FALSE - | _ \ AND FALSE (eupdate ex2)) - | _ \ AND (eupdate ex1) (eupdate ex2))" -| "eupdate (OR ex1 ex2) = - (case (eupdate ex1) of - TRUE \ (case (eupdate ex2) of - TRUE \ TRUE - | FALSE \ TRUE - | _ \ OR TRUE (eupdate ex2)) - | FALSE \ (case (eupdate ex2) of - TRUE \ TRUE - | FALSE \ FALSE - | _ \ OR FALSE (eupdate ex2)) - | _ \ OR (eupdate ex1) (eupdate ex2))" -| "eupdate (NOT ex1) = - (case (eupdate ex1) of - TRUE \ FALSE - | FALSE \ TRUE - | _ \ NOT (eupdate ex1))" -| "eupdate (CALL i xs) = CALL i xs" -| "eupdate (ECALL e i xs r) = ECALL e i xs r" - -value "eupdate (UINT 8 250)" -lemma "eupdate (UINT 8 250) - =UINT 8 250" - by(simp) -lemma "eupdate (UINT 8 500) - = UINT 8 244" - by(simp) -lemma "eupdate (E.INT 8 (-100)) - = E.INT 8 (- 100)" - by(simp) -lemma "eupdate (E.INT 8 (-150)) - = E.INT 8 106" - by(simp) -lemma "eupdate (PLUS (UINT 8 100) (UINT 8 100)) - = UINT 8 200" - by(simp) -lemma "eupdate (PLUS (UINT 8 257) (UINT 16 100)) - = UINT 16 101" - by(simp) -lemma "eupdate (PLUS (E.INT 8 100) (UINT 8 250)) - = PLUS (E.INT 8 100) (UINT 8 250)" - by(simp) -lemma "eupdate (PLUS (E.INT 8 250) (UINT 8 500)) - = PLUS (E.INT 8 (- 6)) (UINT 8 244)" - by(simp) -lemma "eupdate (PLUS (E.INT 16 250) (UINT 8 500)) - = E.INT 16 494" - by(simp) -lemma "eupdate (EQUAL (UINT 16 250) (UINT 8 250)) - = TRUE" - by(simp) -lemma "eupdate (EQUAL (E.INT 16 100) (UINT 8 100)) - = TRUE" - by(simp) -lemma "eupdate (EQUAL (E.INT 8 100) (UINT 8 100)) - = EQUAL (E.INT 8 100) (UINT 8 100)" - by(simp) - -lemma update_bounds_int: - assumes "eupdate ex = (E.INT b v)" and "b\vbits" - shows "(v < 2^(b-1)) \ v \ -(2^(b-1))" -proof (cases ex) - case (INT b' v') - then show ?thesis - proof cases - assume "b'\vbits" - show ?thesis - proof cases - let ?x="-(2^(b'-1)) + (v'+2^(b'-1)) mod 2^b'" - assume "v'\0" - with `b'\vbits` have "eupdate (E.INT b' v') = E.INT b' ?x" by simp - with assms have "b=b'" and "v=?x" using INT by (simp,simp) - moreover from `b'\vbits` have "b'>0" by auto - hence "?x < 2 ^(b'-1)" using upper_bound2[of b' "(v' + 2 ^ (b' - 1)) mod 2^b'"] by simp - moreover have "?x \ -(2^(b'-1))" by simp - ultimately show ?thesis by simp - next - let ?x="2^(b'-1) - (-v'+2^(b'-1)-1) mod (2^b') - 1" - assume "\v'\0" - with `b'\vbits` have "eupdate (E.INT b' v') = E.INT b' ?x" by simp - with assms have "b=b'" and "v=?x" using INT by (simp,simp) - moreover have "(-v'+2^(b'-1)-1) mod (2^b')\0" by simp - hence "?x < 2 ^(b'-1)" by arith - moreover from `b'\vbits` have "b'>0" by auto - hence "?x \ -(2^(b'-1))" using lower_bound2[of b' v'] by simp - ultimately show ?thesis by simp - qed - next - assume "\ b'\vbits" - with assms show ?thesis using INT by simp - qed -next - case (UINT b' v') - with assms show ?thesis - proof cases - assume "b'\vbits" - with assms show ?thesis using UINT by simp - next - assume "\ b'\vbits" - with assms show ?thesis using UINT by simp - qed -next - case (ADDRESS x3) - with assms show ?thesis by simp -next - case (BALANCE x4) - with assms show ?thesis by simp -next - case THIS - with assms show ?thesis by simp -next - case SENDER - with assms show ?thesis by simp -next - case VALUE - with assms show ?thesis by simp -next - case TRUE - with assms show ?thesis by simp -next - case FALSE - with assms show ?thesis by simp -next - case (LVAL x7) - with assms show ?thesis by simp -next - case p: (PLUS e1 e2) - show ?thesis - proof (cases "eupdate e1") - case i: (INT b1 v1) - show ?thesis - proof cases - assume "b1\vbits" - show ?thesis - proof (cases "eupdate e2") - case i2: (INT b2 v2) - then show ?thesis - proof cases - let ?v="v1+v2" - assume "b2\vbits" - show ?thesis - proof cases - let ?x="-(2^((max b1 b2)-1)) + (?v+2^((max b1 b2)-1)) mod 2^(max b1 b2)" - assume "?v\0" - with `b1\vbits` `b2\vbits` i i2 have "eupdate (PLUS e1 e2) = E.INT (max b1 b2) ?x" by simp - with assms have "b=max b1 b2" and "v=?x" using p by (simp,simp) - moreover from `b1\vbits` have "max b1 b2>0" by auto - hence "?x < 2 ^(max b1 b2 - 1)" - using upper_bound2[of "max b1 b2" "(?v + 2 ^ (max b1 b2 - 1)) mod 2^max b1 b2"] by simp - moreover have "?x \ -(2^(max b1 b2-1))" by simp - ultimately show ?thesis by simp - next - let ?x="2^((max b1 b2)-1) - (-?v+2^((max b1 b2)-1)-1) mod (2^(max b1 b2)) - 1" - assume "\?v\0" - with `b1\vbits` `b2\vbits` i i2 have "eupdate (PLUS e1 e2) = E.INT (max b1 b2) ?x" by simp - with assms have "b=max b1 b2" and "v=?x" using p by (simp,simp) - moreover have "(-?v+2^(max b1 b2-1)-1) mod (2^max b1 b2)\0" by simp - hence "?x < 2 ^(max b1 b2-1)" by arith - moreover from `b1\vbits` have "max b1 b2>0" by auto - hence "?x \ -(2^(max b1 b2-1))" using lower_bound2[of "max b1 b2" ?v] by simp - ultimately show ?thesis by simp - qed - next - assume "b2\vbits" - with p i i2 `b1\vbits` show ?thesis using assms by simp - qed - next - case u: (UINT b2 v2) - then show ?thesis - proof cases - let ?v="v1+v2" - assume "b2\vbits" - show ?thesis - proof cases - assume "b20" - with `b1\vbits` `b2\vbits` `b2vbits` have "b1>0" by auto - hence "?x < 2 ^(b1 - 1)" using upper_bound2[of b1] by simp - moreover have "?x \ -(2^(b1-1))" by simp - ultimately show ?thesis by simp - next - let ?x="2^(b1-1) - (-?v+2^(b1-1)-1) mod (2^b1) - 1" - assume "\?v\0" - with `b1\vbits` `b2\vbits` `b20" by simp - hence "?x < 2 ^(b1-1)" by arith - moreover from `b1\vbits` have "b1>0" by auto - hence "?x \ -(2^(b1-1))" using lower_bound2[of b1 ?v] by simp - ultimately show ?thesis by simp - qed - next - assume "\ b2vbits` show ?thesis using assms by simp - qed - next - assume "b2\vbits" - with p i u `b1\vbits` show ?thesis using assms by simp - qed - next - case (ADDRESS x3) - with p i `b1\vbits` show ?thesis using assms by simp - next - case (BALANCE x4) - with p i `b1\vbits` show ?thesis using assms by simp - next - case THIS - with p i `b1\vbits` show ?thesis using assms by simp - next - case SENDER - with p i `b1\vbits` show ?thesis using assms by simp - next - case VALUE - with p i `b1\vbits` show ?thesis using assms by simp - next - case TRUE - with p i `b1\vbits` show ?thesis using assms by simp - next - case FALSE - with p i `b1\vbits` show ?thesis using assms by simp - next - case (LVAL x7) - with p i `b1\vbits` show ?thesis using assms by simp - next - case (PLUS x81 x82) - with p i `b1\vbits` show ?thesis using assms by simp - next - case (MINUS x91 x92) - with p i `b1\vbits` show ?thesis using assms by simp - next - case (EQUAL x101 x102) - with p i `b1\vbits` show ?thesis using assms by simp - next - case (LESS x111 x112) - with p i `b1\vbits` show ?thesis using assms by simp - next - case (AND x121 x122) - with p i `b1\vbits` show ?thesis using assms by simp - next - case (OR x131 x132) - with p i `b1\vbits` show ?thesis using assms by simp - next - case (NOT x131) - with p i `b1\vbits` show ?thesis using assms by simp - next - case (CALL x181 x182) - with p i `b1\vbits` show ?thesis using assms by simp - next - case (ECALL x191 x192 x193 x194) - with p i `b1\vbits` show ?thesis using assms by simp - qed - next - assume "\ b1\vbits" - with p i show ?thesis using assms by simp - qed - next - case u: (UINT b1 v1) - show ?thesis - proof cases - assume "b1\vbits" - show ?thesis - proof (cases "eupdate e2") - case i: (INT b2 v2) - then show ?thesis - proof cases - let ?v="v1+v2" - assume "b2\vbits" - show ?thesis - proof cases - assume "b10" - with `b1\vbits` `b2\vbits` `b1vbits` have "b2>0" by auto - hence "?x < 2 ^(b2 - 1)" using upper_bound2[of b2] by simp - moreover have "?x \ -(2^(b2-1))" by simp - ultimately show ?thesis by simp - next - let ?x="2^(b2-1) - (-?v+2^(b2-1)-1) mod (2^b2) - 1" - assume "\?v\0" - with `b1\vbits` `b2\vbits` `b10" by simp - hence "?x < 2 ^(b2-1)" by arith - moreover from `b2\vbits` have "b2>0" by auto - hence "?x \ -(2^(b2-1))" using lower_bound2[of b2 ?v] by simp - ultimately show ?thesis by simp - qed - next - assume "\ b1vbits` show ?thesis using assms by simp - qed - next - assume "b2\vbits" - with p i u `b1\vbits` show ?thesis using assms by simp - qed - next - case u2: (UINT b2 v2) - then show ?thesis - proof cases - assume "b2\vbits" - with `b1\vbits` u u2 p show ?thesis using assms by simp - next - assume "\b2\vbits" - with p u u2 `b1\vbits` show ?thesis using assms by simp - qed - next - case (ADDRESS x3) - with p u `b1\vbits` show ?thesis using assms by simp - next - case (BALANCE x4) - with p u `b1\vbits` show ?thesis using assms by simp - next - case THIS - with p u `b1\vbits` show ?thesis using assms by simp - next - case SENDER - with p u `b1\vbits` show ?thesis using assms by simp - next - case VALUE - with p u `b1\vbits` show ?thesis using assms by simp - next - case TRUE - with p u `b1\vbits` show ?thesis using assms by simp - next - case FALSE - with p u `b1\vbits` show ?thesis using assms by simp - next - case (LVAL x7) - with p u `b1\vbits` show ?thesis using assms by simp - next - case (PLUS x81 x82) - with p u `b1\vbits` show ?thesis using assms by simp - next - case (MINUS x91 x92) - with p u `b1\vbits` show ?thesis using assms by simp - next - case (EQUAL x101 x102) - with p u `b1\vbits` show ?thesis using assms by simp - next - case (LESS x111 x112) - with p u `b1\vbits` show ?thesis using assms by simp - next - case (AND x121 x122) - with p u `b1\vbits` show ?thesis using assms by simp - next - case (OR x131 x132) - with p u `b1\vbits` show ?thesis using assms by simp - next - case (NOT x131) - with p u `b1\vbits` show ?thesis using assms by simp - next - case (CALL x181 x182) - with p u `b1\vbits` show ?thesis using assms by simp - next - case (ECALL x191 x192 x193 x194) - with p u `b1\vbits` show ?thesis using assms by simp - qed - next - assume "\ b1\vbits" - with p u show ?thesis using assms by simp - qed - next - case (ADDRESS x3) - with p show ?thesis using assms by simp - next - case (BALANCE x4) - with p show ?thesis using assms by simp - next - case THIS - with p show ?thesis using assms by simp - next - case SENDER - with p show ?thesis using assms by simp - next - case VALUE - with p show ?thesis using assms by simp - next - case TRUE - with p show ?thesis using assms by simp - next - case FALSE - with p show ?thesis using assms by simp - next - case (LVAL x7) - with p show ?thesis using assms by simp - next - case (PLUS x81 x82) - with p show ?thesis using assms by simp - next - case (MINUS x91 x92) - with p show ?thesis using assms by simp - next - case (EQUAL x101 x102) - with p show ?thesis using assms by simp - next - case (LESS x111 x112) - with p show ?thesis using assms by simp - next - case (AND x121 x122) - with p show ?thesis using assms by simp - next - case (OR x131 x132) - with p show ?thesis using assms by simp - next - case (NOT x131) - with p show ?thesis using assms by simp - next - case (CALL x181 x182) - with p show ?thesis using assms by simp - next - case (ECALL x191 x192 x193 x194) - with p show ?thesis using assms by simp - qed -next - case m: (MINUS e1 e2) - show ?thesis - proof (cases "eupdate e1") - case i: (INT b1 v1) - with m show ?thesis - proof cases - assume "b1\vbits" - show ?thesis - proof (cases "eupdate e2") - case i2: (INT b2 v2) - then show ?thesis - proof cases - let ?v="v1-v2" - assume "b2\vbits" - with `b1 \ vbits` have - u_def: "eupdate (MINUS e1 e2) = - (let v = v1 - v2 - in if 0 \ v - then E.INT (max b1 b2) - (- (2 ^ (max b1 b2 - 1)) + (v + 2 ^ (max b1 b2 - 1)) mod 2 ^ max b1 b2) - else E.INT (max b1 b2) - (2 ^ (max b1 b2 - 1) - (- v + 2 ^ (max b1 b2 - 1) - 1) mod 2 ^ max b1 b2 - 1))" - using i i2 eupdate.simps(11)[of e1 e2] by simp - show ?thesis - proof cases - let ?x="-(2^((max b1 b2)-1)) + (?v+2^((max b1 b2)-1)) mod 2^(max b1 b2)" - assume "?v\0" - with u_def have "eupdate (MINUS e1 e2) = E.INT (max b1 b2) ?x" by simp - with assms have "b=max b1 b2" and "v=?x" using m by (simp,simp) - moreover from `b1\vbits` have "max b1 b2>0" by auto - hence "?x < 2 ^(max b1 b2 - 1)" - using upper_bound2[of "max b1 b2" "(?v + 2 ^ (max b1 b2 - 1)) mod 2^max b1 b2"] by simp - moreover have "?x \ -(2^(max b1 b2-1))" by simp - ultimately show ?thesis by simp - next - let ?x="2^((max b1 b2)-1) - (-?v+2^((max b1 b2)-1)-1) mod (2^(max b1 b2)) - 1" - assume "\?v\0" - with u_def have "eupdate (MINUS e1 e2) = E.INT (max b1 b2) ?x" using u_def by simp - with assms have "b=max b1 b2" and "v=?x" using m by (simp,simp) - moreover have "(-?v+2^(max b1 b2-1)-1) mod (2^max b1 b2)\0" by simp - hence "?x < 2 ^(max b1 b2-1)" by arith - moreover from `b1\vbits` have "max b1 b2>0" by auto - hence "?x \ -(2^(max b1 b2-1))" using lower_bound2[of "max b1 b2" ?v] by simp - ultimately show ?thesis by simp - qed - next - assume "b2\vbits" - with m i i2 `b1\vbits` show ?thesis using assms by simp - qed - next - case u: (UINT b2 v2) - then show ?thesis - proof cases - let ?v="v1-v2" - assume "b2\vbits" - show ?thesis - proof cases - assume "b2 vbits` `b2 \ vbits` have - u_def: "eupdate (MINUS e1 e2) = - (let v = v1 - v2 - in if 0 \ v - then E.INT b1 (- (2 ^ (b1 - 1)) + (v + 2 ^ (b1 - 1)) mod 2 ^ b1) - else E.INT b1 (2 ^ (b1 - 1) - (- v + 2 ^ (b1 - 1) - 1) mod 2 ^ b1 - 1))" - using i u eupdate.simps(11)[of e1 e2] by simp - then show ?thesis - proof cases - let ?x="(-(2^(b1-1)) + (?v+2^(b1-1)) mod (2^b1))" - assume "?v\0" - with u_def have "eupdate (MINUS e1 e2) = E.INT b1 ?x" by simp - with assms have "b=b1" and "v=?x" using m by (simp,simp) - moreover from `b1\vbits` have "b1>0" by auto - hence "?x < 2 ^(b1 - 1)" using upper_bound2[of b1] by simp - moreover have "?x \ -(2^(b1-1))" by simp - ultimately show ?thesis by simp - next - let ?x="2^(b1-1) - (-?v+2^(b1-1)-1) mod (2^b1) - 1" - assume "\?v\0" - with u_def have "eupdate (MINUS e1 e2) = E.INT b1 ?x" by simp - with assms have "b=b1" and "v=?x" using m i u by (simp,simp) - moreover have "(-?v+2^(b1-1)-1) mod 2^b1\0" by simp - hence "?x < 2 ^(b1-1)" by arith - moreover from `b1\vbits` have "b1>0" by auto - hence "?x \ -(2^(b1-1))" using lower_bound2[of b1 ?v] by simp - ultimately show ?thesis by simp - qed - next - assume "\ b2vbits` show ?thesis using assms by simp - qed - next - assume "b2\vbits" - with m i u `b1\vbits` show ?thesis using assms by simp - qed - next - case (ADDRESS x3) - with m i `b1\vbits` show ?thesis using assms by simp - next - case (BALANCE x4) - with m i `b1\vbits` show ?thesis using assms by simp - next - case THIS - with m i `b1\vbits` show ?thesis using assms by simp - next - case SENDER - with m i `b1\vbits` show ?thesis using assms by simp - next - case VALUE - with m i `b1\vbits` show ?thesis using assms by simp - next - case TRUE - with m i `b1\vbits` show ?thesis using assms by simp - next - case FALSE - with m i `b1\vbits` show ?thesis using assms by simp - next - case (LVAL x7) - with m i `b1\vbits` show ?thesis using assms by simp - next - case (PLUS x81 x82) - with m i `b1\vbits` show ?thesis using assms by simp - next - case (MINUS x91 x92) - with m i `b1\vbits` show ?thesis using assms by simp - next - case (EQUAL x101 x102) - with m i `b1\vbits` show ?thesis using assms by simp - next - case (LESS x111 x112) - with m i `b1\vbits` show ?thesis using assms by simp - next - case (AND x121 x122) - with m i `b1\vbits` show ?thesis using assms by simp - next - case (OR x131 x132) - with m i `b1\vbits` show ?thesis using assms by simp - next - case (NOT x131) - with m i `b1\vbits` show ?thesis using assms by simp - next - case (CALL x181 x182) - with m i `b1\vbits` show ?thesis using assms by simp - next - case (ECALL x191 x192 x193 x194) - with m i `b1\vbits` show ?thesis using assms by simp - qed - next - assume "\ b1\vbits" - with m i show ?thesis using assms by simp - qed - next - case u: (UINT b1 v1) - show ?thesis - proof cases - assume "b1\vbits" - show ?thesis - proof (cases "eupdate e2") - case i: (INT b2 v2) - then show ?thesis - proof cases - let ?v="v1-v2" - assume "b2\vbits" - show ?thesis - proof cases - assume "b1 vbits` `b2 \ vbits` have - u_def: "eupdate (MINUS e1 e2) = - (let v = v1 - v2 - in if 0 \ v - then E.INT b2 (- (2 ^ (b2 - 1)) + (v + 2 ^ (b2 - 1)) mod 2 ^ b2) - else E.INT b2 (2 ^ (b2 - 1) - (- v + 2 ^ (b2 - 1) - 1) mod 2 ^ b2 - 1))" - using i u eupdate.simps(11)[of e1 e2] by simp - then show ?thesis - proof cases - let ?x="(-(2^(b2-1)) + (?v+2^(b2-1)) mod (2^b2))" - assume "?v\0" - with u_def have "eupdate (MINUS e1 e2) = E.INT b2 ?x" by simp - with assms have "b=b2" and "v=?x" using m by (simp,simp) - moreover from `b2\vbits` have "b2>0" by auto - hence "?x < 2 ^(b2 - 1)" using upper_bound2[of b2] by simp - moreover have "?x \ -(2^(b2-1))" by simp - ultimately show ?thesis by simp - next - let ?x="2^(b2-1) - (-?v+2^(b2-1)-1) mod (2^b2) - 1" - assume "\?v\0" - with u_def have "eupdate (MINUS e1 e2) = E.INT b2 ?x" by simp - with assms have "b=b2" and "v=?x" using m i u by (simp,simp) - moreover have "(-?v+2^(b2-1)-1) mod 2^b2\0" by simp - hence "?x < 2 ^(b2-1)" by arith - moreover from `b2\vbits` have "b2>0" by auto - hence "?x \ -(2^(b2-1))" using lower_bound2[of b2 ?v] by simp - ultimately show ?thesis by simp - qed - next - assume "\ b1vbits` show ?thesis using assms by simp - qed - next - assume "b2\vbits" - with m i u `b1\vbits` show ?thesis using assms by simp - qed - next - case u2: (UINT b2 v2) - then show ?thesis - proof cases - assume "b2\vbits" - with `b1\vbits` u u2 m show ?thesis using assms by simp - next - assume "\b2\vbits" - with m u u2 `b1\vbits` show ?thesis using assms by simp - qed - next - case (ADDRESS x3) - with m u `b1\vbits` show ?thesis using assms by simp - next - case (BALANCE x4) - with m u `b1\vbits` show ?thesis using assms by simp - next - case THIS - with m u `b1\vbits` show ?thesis using assms by simp - next - case SENDER - with m u `b1\vbits` show ?thesis using assms by simp - next - case VALUE - with m u `b1\vbits` show ?thesis using assms by simp - next - case TRUE - with m u `b1\vbits` show ?thesis using assms by simp - next - case FALSE - with m u `b1\vbits` show ?thesis using assms by simp - next - case (LVAL x7) - with m u `b1\vbits` show ?thesis using assms by simp - next - case (PLUS x81 x82) - with m u `b1\vbits` show ?thesis using assms by simp - next - case (MINUS x91 x92) - with m u `b1\vbits` show ?thesis using assms by simp - next - case (EQUAL x101 x102) - with m u `b1\vbits` show ?thesis using assms by simp - next - case (LESS x111 x112) - with m u `b1\vbits` show ?thesis using assms by simp - next - case (AND x121 x122) - with m u `b1\vbits` show ?thesis using assms by simp - next - case (OR x131 x132) - with m u `b1\vbits` show ?thesis using assms by simp - next - case (NOT x131) - with m u `b1\vbits` show ?thesis using assms by simp - next - case (CALL x181 x182) - with m u `b1\vbits` show ?thesis using assms by simp - next - case (ECALL x191 x192 x193 x194) - with m u `b1\vbits` show ?thesis using assms by simp - qed - next - assume "\ b1\vbits" - with m u show ?thesis using assms by simp - qed - next - case (ADDRESS x3) - with m show ?thesis using assms by simp - next - case (BALANCE x4) - with m show ?thesis using assms by simp - next - case THIS - with m show ?thesis using assms by simp - next - case SENDER - with m show ?thesis using assms by simp - next - case VALUE - with m show ?thesis using assms by simp - next - case TRUE - with m show ?thesis using assms by simp - next - case FALSE - with m show ?thesis using assms by simp - next - case (LVAL x7) - with m show ?thesis using assms by simp - next - case (PLUS x81 x82) - with m show ?thesis using assms by simp - next - case (MINUS x91 x92) - with m show ?thesis using assms by simp - next - case (EQUAL x101 x102) - with m show ?thesis using assms by simp - next - case (LESS x111 x112) - with m show ?thesis using assms by simp - next - case (AND x121 x122) - with m show ?thesis using assms by simp - next - case (OR x131 x132) - with m show ?thesis using assms by simp - next - case (NOT x131) - with m show ?thesis using assms by simp - next - case (CALL x181 x182) - with m show ?thesis using assms by simp - next - case (ECALL x191 x192 x193 x194) - with m show ?thesis using assms by simp - qed -next - case e: (EQUAL e1 e2) - show ?thesis - proof (cases "eupdate e1") - case i: (INT b1 v1) - show ?thesis - proof cases - assume "b1\vbits" - show ?thesis - proof (cases "eupdate e2") - case i2: (INT b2 v2) - then show ?thesis - proof cases - assume "b2\vbits" - show ?thesis - proof cases - assume "v1=v2" - with assms show ?thesis using e i i2 `b1\vbits` `b2\vbits` by simp - next - assume "\ v1=v2" - with assms show ?thesis using e i i2 `b1\vbits` `b2\vbits` by simp - qed - next - assume "b2\vbits" - with e i i2 `b1\vbits` show ?thesis using assms by simp - qed - next - case u: (UINT b2 v2) - then show ?thesis - proof cases - assume "b2\vbits" - show ?thesis - proof cases - assume "b2vbits` `b2\vbits` `b2 v1=v2" - with assms show ?thesis using e i u `b1\vbits` `b2\vbits` `b2 b2vbits` show ?thesis using assms by simp - qed - next - assume "b2\vbits" - with e i u `b1\vbits` show ?thesis using assms by simp - qed - next - case (ADDRESS x3) - with e i `b1\vbits` show ?thesis using assms by simp - next - case (BALANCE x4) - with e i `b1\vbits` show ?thesis using assms by simp - next - case THIS - with e i `b1\vbits` show ?thesis using assms by simp - next - case SENDER - with e i `b1\vbits` show ?thesis using assms by simp - next - case VALUE - with e i `b1\vbits` show ?thesis using assms by simp - next - case TRUE - with e i `b1\vbits` show ?thesis using assms by simp - next - case FALSE - with e i `b1\vbits` show ?thesis using assms by simp - next - case (LVAL x7) - with e i `b1\vbits` show ?thesis using assms by simp - next - case (PLUS x81 x82) - with e i `b1\vbits` show ?thesis using assms by simp - next - case (MINUS x91 x92) - with e i `b1\vbits` show ?thesis using assms by simp - next - case (EQUAL x101 x102) - with e i `b1\vbits` show ?thesis using assms by simp - next - case (LESS x111 x112) - with e i `b1\vbits` show ?thesis using assms by simp - next - case (AND x121 x122) - with e i `b1\vbits` show ?thesis using assms by simp - next - case (OR x131 x132) - with e i `b1\vbits` show ?thesis using assms by simp - next - case (NOT x131) - with e i `b1\vbits` show ?thesis using assms by simp - next - case (CALL x181 x182) - with e i `b1\vbits` show ?thesis using assms by simp - next - case (ECALL x191 x192 x193 x194) - with e i `b1\vbits` show ?thesis using assms by simp - qed - next - assume "\ b1\vbits" - with e i show ?thesis using assms by simp - qed - next - case u: (UINT b1 v1) - show ?thesis - proof cases - assume "b1\vbits" - show ?thesis - proof (cases "eupdate e2") - case i: (INT b2 v2) - then show ?thesis - proof cases - assume "b2\vbits" - show ?thesis - proof cases - assume "b1vbits` `b2\vbits` `b1 v1=v2" - with assms show ?thesis using e i u `b1\vbits` `b2\vbits` `b1 b1vbits` show ?thesis using assms by simp - qed - next - assume "b2\vbits" - with e i u `b1\vbits` show ?thesis using assms by simp - qed - next - case u2: (UINT b2 v2) - then show ?thesis - proof cases - assume "b2\vbits" - show ?thesis - proof cases - assume "v1=v2" - with assms show ?thesis using e u u2 `b1\vbits` `b2\vbits` by simp - next - assume "\ v1=v2" - with assms show ?thesis using e u u2 `b1\vbits` `b2\vbits` by simp - qed - next - assume "\b2\vbits" - with e u u2 `b1\vbits` show ?thesis using assms by simp - qed - next - case (ADDRESS x3) - with e u `b1\vbits` show ?thesis using assms by simp - next - case (BALANCE x4) - with e u `b1\vbits` show ?thesis using assms by simp - next - case THIS - with e u `b1\vbits` show ?thesis using assms by simp - next - case SENDER - with e u `b1\vbits` show ?thesis using assms by simp - next - case VALUE - with e u `b1\vbits` show ?thesis using assms by simp - next - case TRUE - with e u `b1\vbits` show ?thesis using assms by simp - next - case FALSE - with e u `b1\vbits` show ?thesis using assms by simp - next - case (LVAL x7) - with e u `b1\vbits` show ?thesis using assms by simp - next - case (PLUS x81 x82) - with e u `b1\vbits` show ?thesis using assms by simp - next - case (MINUS x91 x92) - with e u `b1\vbits` show ?thesis using assms by simp - next - case (EQUAL x101 x102) - with e u `b1\vbits` show ?thesis using assms by simp - next - case (LESS x111 x112) - with e u `b1\vbits` show ?thesis using assms by simp - next - case (AND x121 x122) - with e u `b1\vbits` show ?thesis using assms by simp - next - case (OR x131 x132) - with e u `b1\vbits` show ?thesis using assms by simp - next - case (NOT x131) - with e u `b1\vbits` show ?thesis using assms by simp - next - case (CALL x181 x182) - with e u `b1\vbits` show ?thesis using assms by simp - next - case (ECALL x191 x192 x193 x194) - with e u `b1\vbits` show ?thesis using assms by simp - qed - next - assume "\ b1\vbits" - with e u show ?thesis using assms by simp - qed - next - case (ADDRESS x3) - with e show ?thesis using assms by simp - next - case (BALANCE x4) - with e show ?thesis using assms by simp - next - case THIS - with e show ?thesis using assms by simp - next - case SENDER - with e show ?thesis using assms by simp - next - case VALUE - with e show ?thesis using assms by simp - next - case TRUE - with e show ?thesis using assms by simp - next - case FALSE - with e show ?thesis using assms by simp - next - case (LVAL x7) - with e show ?thesis using assms by simp - next - case (PLUS x81 x82) - with e show ?thesis using assms by simp - next - case (MINUS x91 x92) - with e show ?thesis using assms by simp - next - case (EQUAL x101 x102) - with e show ?thesis using assms by simp - next - case (LESS x111 x112) - with e show ?thesis using assms by simp - next - case (AND x121 x122) - with e show ?thesis using assms by simp - next - case (OR x131 x132) - with e show ?thesis using assms by simp - next - case (NOT x131) - with e show ?thesis using assms by simp - next - case (CALL x181 x182) - with e show ?thesis using assms by simp - next - case (ECALL x191 x192 x193 x194) - with e show ?thesis using assms by simp - qed -next - case l: (LESS e1 e2) - show ?thesis - proof (cases "eupdate e1") - case i: (INT b1 v1) - show ?thesis - proof cases - assume "b1\vbits" - show ?thesis - proof (cases "eupdate e2") - case i2: (INT b2 v2) - then show ?thesis - proof cases - assume "b2\vbits" - show ?thesis - proof cases - assume "v1vbits` `b2\vbits` by simp - next - assume "\ v1vbits` `b2\vbits` by simp - qed - next - assume "b2\vbits" - with l i i2 `b1\vbits` show ?thesis using assms by simp - qed - next - case u: (UINT b2 v2) - then show ?thesis - proof cases - assume "b2\vbits" - show ?thesis - proof cases - assume "b2vbits` `b2\vbits` `b2 v1vbits` `b2\vbits` `b2 b2vbits` show ?thesis using assms by simp - qed - next - assume "b2\vbits" - with l i u `b1\vbits` show ?thesis using assms by simp - qed - next - case (ADDRESS x3) - with l i `b1\vbits` show ?thesis using assms by simp - next - case (BALANCE x4) - with l i `b1\vbits` show ?thesis using assms by simp - next - case THIS - with l i `b1\vbits` show ?thesis using assms by simp - next - case SENDER - with l i `b1\vbits` show ?thesis using assms by simp - next - case VALUE - with l i `b1\vbits` show ?thesis using assms by simp - next - case TRUE - with l i `b1\vbits` show ?thesis using assms by simp - next - case FALSE - with l i `b1\vbits` show ?thesis using assms by simp - next - case (LVAL x7) - with l i `b1\vbits` show ?thesis using assms by simp - next - case (PLUS x81 x82) - with l i `b1\vbits` show ?thesis using assms by simp - next - case (MINUS x91 x92) - with l i `b1\vbits` show ?thesis using assms by simp - next - case (EQUAL x101 x102) - with l i `b1\vbits` show ?thesis using assms by simp - next - case (LESS x111 x112) - with l i `b1\vbits` show ?thesis using assms by simp - next - case (AND x121 x122) - with l i `b1\vbits` show ?thesis using assms by simp - next - case (OR x131 x132) - with l i `b1\vbits` show ?thesis using assms by simp - next - case (NOT x131) - with l i `b1\vbits` show ?thesis using assms by simp - next - case (CALL x181 x182) - with l i `b1\vbits` show ?thesis using assms by simp - next - case (ECALL x191 x192 x193 x194) - with l i `b1\vbits` show ?thesis using assms by simp - qed - next - assume "\ b1\vbits" - with l i show ?thesis using assms by simp - qed - next - case u: (UINT b1 v1) - show ?thesis - proof cases - assume "b1\vbits" - show ?thesis - proof (cases "eupdate e2") - case i: (INT b2 v2) - then show ?thesis - proof cases - assume "b2\vbits" - show ?thesis - proof cases - assume "b1vbits` `b2\vbits` `b1 v1vbits` `b2\vbits` `b1 b1vbits` show ?thesis using assms by simp - qed - next - assume "b2\vbits" - with l i u `b1\vbits` show ?thesis using assms by simp - qed - next - case u2: (UINT b2 v2) - then show ?thesis - proof cases - assume "b2\vbits" - show ?thesis - proof cases - assume "v1vbits` `b2\vbits` by simp - next - assume "\ v1vbits` `b2\vbits` by simp - qed - next - assume "\b2\vbits" - with l u u2 `b1\vbits` show ?thesis using assms by simp - qed - next - case (ADDRESS x3) - with l u `b1\vbits` show ?thesis using assms by simp - next - case (BALANCE x4) - with l u `b1\vbits` show ?thesis using assms by simp - next - case THIS - with l u `b1\vbits` show ?thesis using assms by simp - next - case SENDER - with l u `b1\vbits` show ?thesis using assms by simp - next - case VALUE - with l u `b1\vbits` show ?thesis using assms by simp - next - case TRUE - with l u `b1\vbits` show ?thesis using assms by simp - next - case FALSE - with l u `b1\vbits` show ?thesis using assms by simp - next - case (LVAL x7) - with l u `b1\vbits` show ?thesis using assms by simp - next - case (PLUS x81 x82) - with l u `b1\vbits` show ?thesis using assms by simp - next - case (MINUS x91 x92) - with l u `b1\vbits` show ?thesis using assms by simp - next - case (EQUAL x101 x102) - with l u `b1\vbits` show ?thesis using assms by simp - next - case (LESS x111 x112) - with l u `b1\vbits` show ?thesis using assms by simp - next - case (AND x121 x122) - with l u `b1\vbits` show ?thesis using assms by simp - next - case (OR x131 x132) - with l u `b1\vbits` show ?thesis using assms by simp - next - case (NOT x131) - with l u `b1\vbits` show ?thesis using assms by simp - next - case (CALL x181 x182) - with l u `b1\vbits` show ?thesis using assms by simp - next - case (ECALL x191 x192 x193 x194) - with l u `b1\vbits` show ?thesis using assms by simp - qed - next - assume "\ b1\vbits" - with l u show ?thesis using assms by simp - qed - next - case (ADDRESS x3) - with l show ?thesis using assms by simp - next - case (BALANCE x4) - with l show ?thesis using assms by simp - next - case THIS - with l show ?thesis using assms by simp - next - case SENDER - with l show ?thesis using assms by simp - next - case VALUE - with l show ?thesis using assms by simp - next - case TRUE - with l show ?thesis using assms by simp - next - case FALSE - with l show ?thesis using assms by simp - next - case (LVAL x7) - with l show ?thesis using assms by simp - next - case (PLUS x81 x82) - with l show ?thesis using assms by simp - next - case (MINUS x91 x92) - with l show ?thesis using assms by simp - next - case (EQUAL x101 x102) - with l show ?thesis using assms by simp - next - case (LESS x111 x112) - with l show ?thesis using assms by simp - next - case (AND x121 x122) - with l show ?thesis using assms by simp - next - case (OR x131 x132) - with l show ?thesis using assms by simp - next - case (NOT x131) - with l show ?thesis using assms by simp - next - case (CALL x181 x182) - with l show ?thesis using assms by simp - next - case (ECALL x191 x192 x193 x194) - with l show ?thesis using assms by simp - qed -next - case a: (AND e1 e2) - show ?thesis - proof (cases "eupdate e1") - case (INT x11 x12) - with a show ?thesis using assms by simp - next - case (UINT x21 x22) - with a show ?thesis using assms by simp - next - case (ADDRESS x3) - with a show ?thesis using assms by simp - next - case (BALANCE x4) - with a show ?thesis using assms by simp - next - case THIS - with a show ?thesis using assms by simp - next - case SENDER - with a show ?thesis using assms by simp - next - case VALUE - with a show ?thesis using assms by simp - next - case t: TRUE - show ?thesis - proof (cases "eupdate e2") - case (INT x11 x12) - with a t show ?thesis using assms by simp - next - case (UINT x21 x22) - with a t show ?thesis using assms by simp - next - case (ADDRESS x3) - with a t show ?thesis using assms by simp - next - case (BALANCE x4) - with a t show ?thesis using assms by simp - next - case THIS - with a t show ?thesis using assms by simp - next - case SENDER - with a t show ?thesis using assms by simp - next - case VALUE - with a t show ?thesis using assms by simp - next - case TRUE - with a t show ?thesis using assms by simp - next - case FALSE - with a t show ?thesis using assms by simp - next - case (LVAL x7) - with a t show ?thesis using assms by simp - next - case (PLUS x81 x82) - with a t show ?thesis using assms by simp - next - case (MINUS x91 x92) - with a t show ?thesis using assms by simp - next - case (EQUAL x101 x102) - with a t show ?thesis using assms by simp - next - case (LESS x111 x112) - with a t show ?thesis using assms by simp - next - case (AND x121 x122) - with a t show ?thesis using assms by simp - next - case (OR x131 x132) - with a t show ?thesis using assms by simp - next - case (NOT x131) - with a t show ?thesis using assms by simp - next - case (CALL x181 x182) - with a t show ?thesis using assms by simp - next - case (ECALL x191 x192 x193 x194) - with a t show ?thesis using assms by simp - qed - next - case f: FALSE - show ?thesis - proof (cases "eupdate e2") - case (INT x11 x12) - with a f show ?thesis using assms by simp - next - case (UINT x21 x22) - with a f show ?thesis using assms by simp - next - case (ADDRESS x3) - with a f show ?thesis using assms by simp - next - case (BALANCE x4) - with a f show ?thesis using assms by simp - next - case THIS - with a f show ?thesis using assms by simp - next - case SENDER - with a f show ?thesis using assms by simp - next - case VALUE - with a f show ?thesis using assms by simp - next - case TRUE - with a f show ?thesis using assms by simp - next - case FALSE - with a f show ?thesis using assms by simp - next - case (LVAL x7) - with a f show ?thesis using assms by simp - next - case (PLUS x81 x82) - with a f show ?thesis using assms by simp - next - case (MINUS x91 x92) - with a f show ?thesis using assms by simp - next - case (EQUAL x101 x102) - with a f show ?thesis using assms by simp - next - case (LESS x111 x112) - with a f show ?thesis using assms by simp - next - case (AND x121 x122) - with a f show ?thesis using assms by simp - next - case (OR x131 x132) - with a f show ?thesis using assms by simp - next - case (NOT x131) - with a f show ?thesis using assms by simp - next - case (CALL x181 x182) - with a f show ?thesis using assms by simp - next - case (ECALL x191 x192 x193 x194) - with a f show ?thesis using assms by simp - qed - next - case (LVAL x7) - with a show ?thesis using assms by simp - next - case (PLUS x81 x82) - with a show ?thesis using assms by simp - next - case (MINUS x91 x92) - with a show ?thesis using assms by simp - next - case (EQUAL x101 x102) - with a show ?thesis using assms by simp - next - case (LESS x111 x112) - with a show ?thesis using assms by simp - next - case (AND x121 x122) - with a show ?thesis using assms by simp - next - case (OR x131 x132) - with a show ?thesis using assms by simp - next - case (NOT x131) - with a show ?thesis using assms by simp - next - case (CALL x181 x182) - with a show ?thesis using assms by simp - next - case (ECALL x191 x192 x193 x194) - with a show ?thesis using assms by simp - qed -next - case o: (OR e1 e2) - show ?thesis - proof (cases "eupdate e1") - case (INT x11 x12) - with o show ?thesis using assms by simp - next - case (UINT x21 x22) - with o show ?thesis using assms by simp - next - case (ADDRESS x3) - with o show ?thesis using assms by simp - next - case (BALANCE x4) - with o show ?thesis using assms by simp - next - case THIS - with o show ?thesis using assms by simp - next - case SENDER - with o show ?thesis using assms by simp - next - case VALUE - with o show ?thesis using assms by simp - next - case t: TRUE - show ?thesis - proof (cases "eupdate e2") - case (INT x11 x12) - with o t show ?thesis using assms by simp - next - case (UINT x21 x22) - with o t show ?thesis using assms by simp - next - case (ADDRESS x3) - with o t show ?thesis using assms by simp - next - case (BALANCE x4) - with o t show ?thesis using assms by simp - next - case THIS - with o t show ?thesis using assms by simp - next - case SENDER - with o t show ?thesis using assms by simp - next - case VALUE - with o t show ?thesis using assms by simp - next - case TRUE - with o t show ?thesis using assms by simp - next - case FALSE - with o t show ?thesis using assms by simp - next - case (LVAL x7) - with o t show ?thesis using assms by simp - next - case (PLUS x81 x82) - with o t show ?thesis using assms by simp - next - case (MINUS x91 x92) - with o t show ?thesis using assms by simp - next - case (EQUAL x101 x102) - with o t show ?thesis using assms by simp - next - case (LESS x111 x112) - with o t show ?thesis using assms by simp - next - case (AND x121 x122) - with o t show ?thesis using assms by simp - next - case (OR x131 x132) - with o t show ?thesis using assms by simp - next - case (NOT x131) - with o t show ?thesis using assms by simp - next - case (CALL x181 x182) - with o t show ?thesis using assms by simp - next - case (ECALL x191 x192 x193 x194) - with o t show ?thesis using assms by simp - qed - next - case f: FALSE - show ?thesis - proof (cases "eupdate e2") - case (INT x11 x12) - with o f show ?thesis using assms by simp - next - case (UINT x21 x22) - with o f show ?thesis using assms by simp - next - case (ADDRESS x3) - with o f show ?thesis using assms by simp - next - case (BALANCE x4) - with o f show ?thesis using assms by simp - next - case THIS - with o f show ?thesis using assms by simp - next - case SENDER - with o f show ?thesis using assms by simp - next - case VALUE - with o f show ?thesis using assms by simp - next - case TRUE - with o f show ?thesis using assms by simp - next - case FALSE - with o f show ?thesis using assms by simp - next - case (LVAL x7) - with o f show ?thesis using assms by simp - next - case (PLUS x81 x82) - with o f show ?thesis using assms by simp - next - case (MINUS x91 x92) - with o f show ?thesis using assms by simp - next - case (EQUAL x101 x102) - with o f show ?thesis using assms by simp - next - case (LESS x111 x112) - with o f show ?thesis using assms by simp - next - case (AND x121 x122) - with o f show ?thesis using assms by simp - next - case (OR x131 x132) - with o f show ?thesis using assms by simp - next - case (NOT x131) - with o f show ?thesis using assms by simp - next - case (CALL x181 x182) - with o f show ?thesis using assms by simp - next - case (ECALL x191 x192 x193 x194) - with o f show ?thesis using assms by simp - qed - next - case (LVAL x7) - with o show ?thesis using assms by simp - next - case (PLUS x81 x82) - with o show ?thesis using assms by simp - next - case (MINUS x91 x92) - with o show ?thesis using assms by simp - next - case (EQUAL x101 x102) - with o show ?thesis using assms by simp - next - case (LESS x111 x112) - with o show ?thesis using assms by simp - next - case (AND x121 x122) - with o show ?thesis using assms by simp - next - case (OR x131 x132) - with o show ?thesis using assms by simp - next - case (NOT x131) - with o show ?thesis using assms by simp - next - case (CALL x181 x182) - with o show ?thesis using assms by simp - next - case (ECALL x191 x192 x193 x194) - with o show ?thesis using assms by simp - qed -next - case o: (NOT e1) - show ?thesis - proof (cases "eupdate e1") - case (INT x11 x12) - with o show ?thesis using assms by simp - next - case (UINT x21 x22) - with o show ?thesis using assms by simp - next - case (ADDRESS x3) - with o show ?thesis using assms by simp - next - case (BALANCE x4) - with o show ?thesis using assms by simp - next - case THIS - with o show ?thesis using assms by simp - next - case SENDER - with o show ?thesis using assms by simp - next - case VALUE - with o show ?thesis using assms by simp - next - case t: TRUE - with o show ?thesis using assms by simp - next - case f: FALSE - with o show ?thesis using assms by simp - next - case (LVAL x7) - with o show ?thesis using assms by simp - next - case (PLUS x81 x82) - with o show ?thesis using assms by simp - next - case (MINUS x91 x92) - with o show ?thesis using assms by simp - next - case (EQUAL x101 x102) - with o show ?thesis using assms by simp - next - case (LESS x111 x112) - with o show ?thesis using assms by simp - next - case (AND x121 x122) - with o show ?thesis using assms by simp - next - case (OR x131 x132) - with o show ?thesis using assms by simp - next - case (NOT x131) - with o show ?thesis using assms by simp - next - case (CALL x181 x182) - with o show ?thesis using assms by simp - next - case (ECALL x191 x192 x193 x194) - with o show ?thesis using assms by simp - qed -next - case (CALL x181 x182) - with assms show ?thesis by simp -next - case (ECALL x191 x192 x193 x194) - with assms show ?thesis by simp -qed - -lemma update_bounds_uint: - assumes "eupdate ex = UINT b v" and "b\vbits" - shows "v < 2^b \ v \ 0" -proof (cases ex) - case (INT b' v') - with assms show ?thesis - proof cases - assume "b'\vbits" - show ?thesis - proof cases - assume "v'\0" - with INT show ?thesis using assms `b'\vbits` by simp - next - assume "\ v'\0" - with INT show ?thesis using assms `b'\vbits` by simp - qed - next - assume "\ b'\vbits" - with INT show ?thesis using assms by simp - qed -next - case (UINT b' v') - then show ?thesis - proof cases - assume "b'\vbits" - with UINT show ?thesis using assms by auto - next - assume "\ b'\vbits" - with UINT show ?thesis using assms by auto - qed -next - case (ADDRESS x3) - with assms show ?thesis by simp -next - case (BALANCE x4) - with assms show ?thesis by simp -next - case THIS - with assms show ?thesis by simp -next - case SENDER - with assms show ?thesis by simp -next - case VALUE - with assms show ?thesis by simp -next - case TRUE - with assms show ?thesis by simp -next - case FALSE - with assms show ?thesis by simp -next - case (LVAL x7) - with assms show ?thesis by simp -next - case p: (PLUS e1 e2) - show ?thesis - proof (cases "eupdate e1") - case i: (INT b1 v1) - with p show ?thesis - proof cases - assume "b1\vbits" - show ?thesis - proof (cases "eupdate e2") - case i2: (INT b2 v2) - then show ?thesis - proof cases - let ?v="v1+v2" - assume "b2\vbits" - show ?thesis - proof cases - assume "?v\0" - with assms show ?thesis using p i i2 `b1\vbits` `b2\vbits` by simp - next - assume "\?v\0" - with assms show ?thesis using p i i2 `b1\vbits` `b2\vbits` by simp - qed - next - assume "b2\vbits" - with p i i2 `b1\vbits` show ?thesis using assms by simp - qed - next - case u: (UINT b2 v2) - then show ?thesis - proof cases - let ?v="v1+v2" - assume "b2\vbits" - show ?thesis - proof cases - assume "b20" - with assms show ?thesis using p i u `b1\vbits` `b2\vbits` `b2?v\0" - with assms show ?thesis using p i u `b1\vbits` `b2\vbits` `b2 b2vbits` show ?thesis using assms by simp - qed - next - assume "b2\vbits" - with p i u `b1\vbits` show ?thesis using assms by simp - qed - next - case (ADDRESS x3) - with p i `b1\vbits` show ?thesis using assms by simp - next - case (BALANCE x4) - with p i `b1\vbits` show ?thesis using assms by simp - next - case THIS - with p i `b1\vbits` show ?thesis using assms by simp - next - case SENDER - with p i `b1\vbits` show ?thesis using assms by simp - next - case VALUE - with p i `b1\vbits` show ?thesis using assms by simp - next - case TRUE - with p i `b1\vbits` show ?thesis using assms by simp - next - case FALSE - with p i `b1\vbits` show ?thesis using assms by simp - next - case (LVAL x7) - with p i `b1\vbits` show ?thesis using assms by simp - next - case (PLUS x81 x82) - with p i `b1\vbits` show ?thesis using assms by simp - next - case (MINUS x91 x92) - with p i `b1\vbits` show ?thesis using assms by simp - next - case (EQUAL x101 x102) - with p i `b1\vbits` show ?thesis using assms by simp - next - case (LESS x111 x112) - with p i `b1\vbits` show ?thesis using assms by simp - next - case (AND x121 x122) - with p i `b1\vbits` show ?thesis using assms by simp - next - case (OR x131 x132) - with p i `b1\vbits` show ?thesis using assms by simp - next - case (NOT x131) - with p i `b1\vbits` show ?thesis using assms by simp - next - case (CALL x181 x182) - with p i `b1\vbits` show ?thesis using assms by simp - next - case (ECALL x191 x192 x193 x194) - with p i `b1\vbits` show ?thesis using assms by simp - qed - next - assume "\ b1\vbits" - with p i show ?thesis using assms by simp - qed - next - case u: (UINT b1 v1) - with p show ?thesis - proof cases - assume "b1\vbits" - show ?thesis - proof (cases "eupdate e2") - case i: (INT b2 v2) - then show ?thesis - proof cases - let ?v="v1+v2" - assume "b2\vbits" - show ?thesis - proof cases - assume "b10" - with assms show ?thesis using p i u `b1\vbits` `b2\vbits` `b1?v\0" - with assms show ?thesis using p i u `b1\vbits` `b2\vbits` `b1 b1vbits` show ?thesis using assms by simp - qed - next - assume "b2\vbits" - with p i u `b1\vbits` show ?thesis using assms by simp - qed - next - case u2: (UINT b2 v2) - then show ?thesis - proof cases - let ?x="((v1 + v2) mod (2^(max b1 b2)))" - assume "b2\vbits" - with `b1\vbits` u u2 have "eupdate (PLUS e1 e2) = UINT (max b1 b2) ?x" by simp - with assms have "b=max b1 b2" and "v=?x" using p by (simp,simp) - moreover from `b1\vbits` have "max b1 b2>0" by auto - hence "?x < 2 ^(max b1 b2)" by simp - moreover have "?x \ 0" by simp - ultimately show ?thesis by simp - next - assume "\b2\vbits" - with p u u2 `b1\vbits` show ?thesis using assms by simp - qed - next - case (ADDRESS x3) - with p u `b1\vbits` show ?thesis using assms by simp - next - case (BALANCE x4) - with p u `b1\vbits` show ?thesis using assms by simp - next - case THIS - with p u `b1\vbits` show ?thesis using assms by simp - next - case SENDER - with p u `b1\vbits` show ?thesis using assms by simp - next - case VALUE - with p u `b1\vbits` show ?thesis using assms by simp - next - case TRUE - with p u `b1\vbits` show ?thesis using assms by simp - next - case FALSE - with p u `b1\vbits` show ?thesis using assms by simp - next - case (LVAL x7) - with p u `b1\vbits` show ?thesis using assms by simp - next - case (PLUS x81 x82) - with p u `b1\vbits` show ?thesis using assms by simp - next - case (MINUS x91 x92) - with p u `b1\vbits` show ?thesis using assms by simp - next - case (EQUAL x101 x102) - with p u `b1\vbits` show ?thesis using assms by simp - next - case (LESS x111 x112) - with p u `b1\vbits` show ?thesis using assms by simp - next - case (AND x121 x122) - with p u `b1\vbits` show ?thesis using assms by simp - next - case (OR x131 x132) - with p u `b1\vbits` show ?thesis using assms by simp - next - case (NOT x131) - with p u `b1\vbits` show ?thesis using assms by simp - next - case (CALL x181 x182) - with p u `b1\vbits` show ?thesis using assms by simp - next - case (ECALL x191 x192 x193 x194) - with p u `b1\vbits` show ?thesis using assms by simp - qed - next - assume "\ b1\vbits" - with p u show ?thesis using assms by simp - qed - next - case (ADDRESS x3) - with p show ?thesis using assms by simp - next - case (BALANCE x4) - with p show ?thesis using assms by simp - next - case THIS - with p show ?thesis using assms by simp - next - case SENDER - with p show ?thesis using assms by simp - next - case VALUE - with p show ?thesis using assms by simp - next - case TRUE - with p show ?thesis using assms by simp - next - case FALSE - with p show ?thesis using assms by simp - next - case (LVAL x7) - with p show ?thesis using assms by simp - next - case (PLUS x81 x82) - with p show ?thesis using assms by simp - next - case (MINUS x91 x92) - with p show ?thesis using assms by simp - next - case (EQUAL x101 x102) - with p show ?thesis using assms by simp - next - case (LESS x111 x112) - with p show ?thesis using assms by simp - next - case (AND x121 x122) - with p show ?thesis using assms by simp - next - case (OR x131 x132) - with p show ?thesis using assms by simp - next - case (NOT x131) - with p show ?thesis using assms by simp - next - case (CALL x181 x182) - with p show ?thesis using assms by simp - next - case (ECALL x191 x192 x193 x194) - with p show ?thesis using assms by simp - qed -next - case m: (MINUS e1 e2) - show ?thesis - proof (cases "eupdate e1") - case i: (INT b1 v1) - with m show ?thesis - proof cases - assume "b1\vbits" - show ?thesis - proof (cases "eupdate e2") - case i2: (INT b2 v2) - then show ?thesis - proof cases - let ?v="v1-v2" - assume "b2\vbits" - show ?thesis - proof cases - assume "?v\0" - with assms show ?thesis using m i i2 `b1\vbits` `b2\vbits` by simp - next - assume "\?v\0" - with assms show ?thesis using m i i2 `b1\vbits` `b2\vbits` by simp - qed - next - assume "b2\vbits" - with m i i2 `b1\vbits` show ?thesis using assms by simp - qed - next - case u: (UINT b2 v2) - then show ?thesis - proof cases - let ?v="v1-v2" - assume "b2\vbits" - show ?thesis - proof cases - assume "b20" - with assms show ?thesis using m i u `b1\vbits` `b2\vbits` `b2?v\0" - with assms show ?thesis using m i u `b1\vbits` `b2\vbits` `b2 b2vbits` show ?thesis using assms by simp - qed - next - assume "b2\vbits" - with m i u `b1\vbits` show ?thesis using assms by simp - qed - next - case (ADDRESS x3) - with m i `b1\vbits` show ?thesis using assms by simp - next - case (BALANCE x4) - with m i `b1\vbits` show ?thesis using assms by simp - next - case THIS - with m i `b1\vbits` show ?thesis using assms by simp - next - case SENDER - with m i `b1\vbits` show ?thesis using assms by simp - next - case VALUE - with m i `b1\vbits` show ?thesis using assms by simp - next - case TRUE - with m i `b1\vbits` show ?thesis using assms by simp - next - case FALSE - with m i `b1\vbits` show ?thesis using assms by simp - next - case (LVAL x7) - with m i `b1\vbits` show ?thesis using assms by simp - next - case (PLUS x81 x82) - with m i `b1\vbits` show ?thesis using assms by simp - next - case (MINUS x91 x92) - with m i `b1\vbits` show ?thesis using assms by simp - next - case (EQUAL x101 x102) - with m i `b1\vbits` show ?thesis using assms by simp - next - case (LESS x111 x112) - with m i `b1\vbits` show ?thesis using assms by simp - next - case (AND x121 x122) - with m i `b1\vbits` show ?thesis using assms by simp - next - case (OR x131 x132) - with m i `b1\vbits` show ?thesis using assms by simp - next - case (NOT x131) - with m i `b1\vbits` show ?thesis using assms by simp - next - case (CALL x181 x182) - with m i `b1\vbits` show ?thesis using assms by simp - next - case (ECALL x191 x192 x193 x194) - with m i `b1\vbits` show ?thesis using assms by simp - qed - next - assume "\ b1\vbits" - with m i show ?thesis using assms by simp - qed - next - case u: (UINT b1 v1) - with m show ?thesis - proof cases - assume "b1\vbits" - show ?thesis - proof (cases "eupdate e2") - case i: (INT b2 v2) - then show ?thesis - proof cases - let ?v="v1-v2" - assume "b2\vbits" - show ?thesis - proof cases - assume "b10" - with assms show ?thesis using m i u `b1\vbits` `b2\vbits` `b1?v\0" - with assms show ?thesis using m i u `b1\vbits` `b2\vbits` `b1 b1vbits` show ?thesis using assms by simp - qed - next - assume "b2\vbits" - with m i u `b1\vbits` show ?thesis using assms by simp - qed - next - case u2: (UINT b2 v2) - then show ?thesis - proof cases - let ?x="((v1 - v2) mod (2^(max b1 b2)))" - assume "b2\vbits" - with `b1\vbits` u u2 have "eupdate (MINUS e1 e2) = UINT (max b1 b2) ?x" by simp - with assms have "b=max b1 b2" and "v=?x" using m by (simp,simp) - moreover from `b1\vbits` have "max b1 b2>0" by auto - hence "?x < 2 ^(max b1 b2)" by simp - moreover have "?x \ 0" by simp - ultimately show ?thesis by simp - next - assume "\b2\vbits" - with m u u2 `b1\vbits` show ?thesis using assms by simp - qed - next - case (ADDRESS x3) - with m u `b1\vbits` show ?thesis using assms by simp - next - case (BALANCE x4) - with m u `b1\vbits` show ?thesis using assms by simp - next - case THIS - with m u `b1\vbits` show ?thesis using assms by simp - next - case SENDER - with m u `b1\vbits` show ?thesis using assms by simp - next - case VALUE - with m u `b1\vbits` show ?thesis using assms by simp - next - case TRUE - with m u `b1\vbits` show ?thesis using assms by simp - next - case FALSE - with m u `b1\vbits` show ?thesis using assms by simp - next - case (LVAL x7) - with m u `b1\vbits` show ?thesis using assms by simp - next - case (PLUS x81 x82) - with m u `b1\vbits` show ?thesis using assms by simp - next - case (MINUS x91 x92) - with m u `b1\vbits` show ?thesis using assms by simp - next - case (EQUAL x101 x102) - with m u `b1\vbits` show ?thesis using assms by simp - next - case (LESS x111 x112) - with m u `b1\vbits` show ?thesis using assms by simp - next - case (AND x121 x122) - with m u `b1\vbits` show ?thesis using assms by simp - next - case (OR x131 x132) - with m u `b1\vbits` show ?thesis using assms by simp - next - case (NOT x131) - with m u `b1\vbits` show ?thesis using assms by simp - next - case (CALL x181 x182) - with m u `b1\vbits` show ?thesis using assms by simp - next - case (ECALL x191 x192 x193 x194) - with m u `b1\vbits` show ?thesis using assms by simp - qed - next - assume "\ b1\vbits" - with m u show ?thesis using assms by simp - qed - next - case (ADDRESS x3) - with m show ?thesis using assms by simp - next - case (BALANCE x4) - with m show ?thesis using assms by simp - next - case THIS - with m show ?thesis using assms by simp - next - case SENDER - with m show ?thesis using assms by simp - next - case VALUE - with m show ?thesis using assms by simp - next - case TRUE - with m show ?thesis using assms by simp - next - case FALSE - with m show ?thesis using assms by simp - next - case (LVAL x7) - with m show ?thesis using assms by simp - next - case (PLUS x81 x82) - with m show ?thesis using assms by simp - next - case (MINUS x91 x92) - with m show ?thesis using assms by simp - next - case (EQUAL x101 x102) - with m show ?thesis using assms by simp - next - case (LESS x111 x112) - with m show ?thesis using assms by simp - next - case (AND x121 x122) - with m show ?thesis using assms by simp - next - case (OR x131 x132) - with m show ?thesis using assms by simp - next - case (NOT x131) - with m show ?thesis using assms by simp - next - case (CALL x181 x182) - with m show ?thesis using assms by simp - next - case (ECALL x191 x192 x193 x194) - with m show ?thesis using assms by simp - qed -next - case e: (EQUAL e1 e2) - show ?thesis - proof (cases "eupdate e1") - case i: (INT b1 v1) - show ?thesis - proof cases - assume "b1\vbits" - show ?thesis - proof (cases "eupdate e2") - case i2: (INT b2 v2) - then show ?thesis - proof cases - assume "b2\vbits" - show ?thesis - proof cases - assume "v1=v2" - with assms show ?thesis using e i i2 `b1\vbits` `b2\vbits` by simp - next - assume "\ v1=v2" - with assms show ?thesis using e i i2 `b1\vbits` `b2\vbits` by simp - qed - next - assume "b2\vbits" - with e i i2 `b1\vbits` show ?thesis using assms by simp - qed - next - case u: (UINT b2 v2) - then show ?thesis - proof cases - assume "b2\vbits" - show ?thesis - proof cases - assume "b2vbits` `b2\vbits` `b2 v1=v2" - with assms show ?thesis using e i u `b1\vbits` `b2\vbits` `b2 b2vbits` show ?thesis using assms by simp - qed - next - assume "b2\vbits" - with e i u `b1\vbits` show ?thesis using assms by simp - qed - next - case (ADDRESS x3) - with e i `b1\vbits` show ?thesis using assms by simp - next - case (BALANCE x4) - with e i `b1\vbits` show ?thesis using assms by simp - next - case THIS - with e i `b1\vbits` show ?thesis using assms by simp - next - case SENDER - with e i `b1\vbits` show ?thesis using assms by simp - next - case VALUE - with e i `b1\vbits` show ?thesis using assms by simp - next - case TRUE - with e i `b1\vbits` show ?thesis using assms by simp - next - case FALSE - with e i `b1\vbits` show ?thesis using assms by simp - next - case (LVAL x7) - with e i `b1\vbits` show ?thesis using assms by simp - next - case (PLUS x81 x82) - with e i `b1\vbits` show ?thesis using assms by simp - next - case (MINUS x91 x92) - with e i `b1\vbits` show ?thesis using assms by simp - next - case (EQUAL x101 x102) - with e i `b1\vbits` show ?thesis using assms by simp - next - case (LESS x111 x112) - with e i `b1\vbits` show ?thesis using assms by simp - next - case (AND x121 x122) - with e i `b1\vbits` show ?thesis using assms by simp - next - case (OR x131 x132) - with e i `b1\vbits` show ?thesis using assms by simp - next - case (NOT x131) - with e i `b1\vbits` show ?thesis using assms by simp - next - case (CALL x181 x182) - with e i `b1\vbits` show ?thesis using assms by simp - next - case (ECALL x191 x192 x193 x194) - with e i `b1\vbits` show ?thesis using assms by simp - qed - next - assume "\ b1\vbits" - with e i show ?thesis using assms by simp - qed - next - case u: (UINT b1 v1) - show ?thesis - proof cases - assume "b1\vbits" - show ?thesis - proof (cases "eupdate e2") - case i: (INT b2 v2) - then show ?thesis - proof cases - let ?v="v1+v2" - assume "b2\vbits" - show ?thesis - proof cases - assume "b1vbits` `b2\vbits` `b1 v1=v2" - with assms show ?thesis using e i u `b1\vbits` `b2\vbits` `b1 b1vbits` show ?thesis using assms by simp - qed - next - assume "b2\vbits" - with e i u `b1\vbits` show ?thesis using assms by simp - qed - next - case u2: (UINT b2 v2) - then show ?thesis - proof cases - assume "b2\vbits" - show ?thesis - proof cases - assume "v1=v2" - with assms show ?thesis using e u u2 `b1\vbits` `b2\vbits` by simp - next - assume "\ v1=v2" - with assms show ?thesis using e u u2 `b1\vbits` `b2\vbits` by simp - qed - next - assume "\b2\vbits" - with e u u2 `b1\vbits` show ?thesis using assms by simp - qed - next - case (ADDRESS x3) - with e u `b1\vbits` show ?thesis using assms by simp - next - case (BALANCE x4) - with e u `b1\vbits` show ?thesis using assms by simp - next - case THIS - with e u `b1\vbits` show ?thesis using assms by simp - next - case SENDER - with e u `b1\vbits` show ?thesis using assms by simp - next - case VALUE - with e u `b1\vbits` show ?thesis using assms by simp - next - case TRUE - with e u `b1\vbits` show ?thesis using assms by simp - next - case FALSE - with e u `b1\vbits` show ?thesis using assms by simp - next - case (LVAL x7) - with e u `b1\vbits` show ?thesis using assms by simp - next - case (PLUS x81 x82) - with e u `b1\vbits` show ?thesis using assms by simp - next - case (MINUS x91 x92) - with e u `b1\vbits` show ?thesis using assms by simp - next - case (EQUAL x101 x102) - with e u `b1\vbits` show ?thesis using assms by simp - next - case (LESS x111 x112) - with e u `b1\vbits` show ?thesis using assms by simp - next - case (AND x121 x122) - with e u `b1\vbits` show ?thesis using assms by simp - next - case (OR x131 x132) - with e u `b1\vbits` show ?thesis using assms by simp - next - case (NOT x131) - with e u `b1\vbits` show ?thesis using assms by simp - next - case (CALL x181 x182) - with e u `b1\vbits` show ?thesis using assms by simp - next - case (ECALL x191 x192 x193 x194) - with e u `b1\vbits` show ?thesis using assms by simp - qed - next - assume "\ b1\vbits" - with e u show ?thesis using assms by simp - qed - next - case (ADDRESS x3) - with e show ?thesis using assms by simp - next - case (BALANCE x4) - with e show ?thesis using assms by simp - next - case THIS - with e show ?thesis using assms by simp - next - case SENDER - with e show ?thesis using assms by simp - next - case VALUE - with e show ?thesis using assms by simp - next - case TRUE - with e show ?thesis using assms by simp - next - case FALSE - with e show ?thesis using assms by simp - next - case (LVAL x7) - with e show ?thesis using assms by simp - next - case (PLUS x81 x82) - with e show ?thesis using assms by simp - next - case (MINUS x91 x92) - with e show ?thesis using assms by simp - next - case (EQUAL x101 x102) - with e show ?thesis using assms by simp - next - case (LESS x111 x112) - with e show ?thesis using assms by simp - next - case (AND x121 x122) - with e show ?thesis using assms by simp - next - case (OR x131 x132) - with e show ?thesis using assms by simp - next - case (NOT x131) - with e show ?thesis using assms by simp - next - case (CALL x181 x182) - with e show ?thesis using assms by simp - next - case (ECALL x191 x192 x193 x194) - with e show ?thesis using assms by simp - qed -next - case l: (LESS e1 e2) - show ?thesis - proof (cases "eupdate e1") - case i: (INT b1 v1) - show ?thesis - proof cases - assume "b1\vbits" - show ?thesis - proof (cases "eupdate e2") - case i2: (INT b2 v2) - then show ?thesis - proof cases - assume "b2\vbits" - show ?thesis - proof cases - assume "v1vbits` `b2\vbits` by simp - next - assume "\ v1vbits` `b2\vbits` by simp - qed - next - assume "b2\vbits" - with l i i2 `b1\vbits` show ?thesis using assms by simp - qed - next - case u: (UINT b2 v2) - then show ?thesis - proof cases - assume "b2\vbits" - show ?thesis - proof cases - assume "b2vbits` `b2\vbits` `b2 v1vbits` `b2\vbits` `b2 b2vbits` show ?thesis using assms by simp - qed - next - assume "b2\vbits" - with l i u `b1\vbits` show ?thesis using assms by simp - qed - next - case (ADDRESS x3) - with l i `b1\vbits` show ?thesis using assms by simp - next - case (BALANCE x4) - with l i `b1\vbits` show ?thesis using assms by simp - next - case THIS - with l i `b1\vbits` show ?thesis using assms by simp - next - case SENDER - with l i `b1\vbits` show ?thesis using assms by simp - next - case VALUE - with l i `b1\vbits` show ?thesis using assms by simp - next - case TRUE - with l i `b1\vbits` show ?thesis using assms by simp - next - case FALSE - with l i `b1\vbits` show ?thesis using assms by simp - next - case (LVAL x7) - with l i `b1\vbits` show ?thesis using assms by simp - next - case (PLUS x81 x82) - with l i `b1\vbits` show ?thesis using assms by simp - next - case (MINUS x91 x92) - with l i `b1\vbits` show ?thesis using assms by simp - next - case (EQUAL x101 x102) - with l i `b1\vbits` show ?thesis using assms by simp - next - case (LESS x111 x112) - with l i `b1\vbits` show ?thesis using assms by simp - next - case (AND x121 x122) - with l i `b1\vbits` show ?thesis using assms by simp - next - case (OR x131 x132) - with l i `b1\vbits` show ?thesis using assms by simp - next - case (NOT x131) - with l i `b1\vbits` show ?thesis using assms by simp - next - case (CALL x181 x182) - with l i `b1\vbits` show ?thesis using assms by simp - next - case (ECALL x191 x192 x193 x194) - with l i `b1\vbits` show ?thesis using assms by simp - qed - next - assume "\ b1\vbits" - with l i show ?thesis using assms by simp - qed - next - case u: (UINT b1 v1) - show ?thesis - proof cases - assume "b1\vbits" - show ?thesis - proof (cases "eupdate e2") - case i: (INT b2 v2) - then show ?thesis - proof cases - let ?v="v1+v2" - assume "b2\vbits" - show ?thesis - proof cases - assume "b1vbits` `b2\vbits` `b1 v1vbits` `b2\vbits` `b1 b1vbits` show ?thesis using assms by simp - qed - next - assume "b2\vbits" - with l i u `b1\vbits` show ?thesis using assms by simp - qed - next - case u2: (UINT b2 v2) - then show ?thesis - proof cases - assume "b2\vbits" - show ?thesis - proof cases - assume "v1vbits` `b2\vbits` by simp - next - assume "\ v1vbits` `b2\vbits` by simp - qed - next - assume "\b2\vbits" - with l u u2 `b1\vbits` show ?thesis using assms by simp - qed - next - case (ADDRESS x3) - with l u `b1\vbits` show ?thesis using assms by simp - next - case (BALANCE x4) - with l u `b1\vbits` show ?thesis using assms by simp - next - case THIS - with l u `b1\vbits` show ?thesis using assms by simp - next - case SENDER - with l u `b1\vbits` show ?thesis using assms by simp - next - case VALUE - with l u `b1\vbits` show ?thesis using assms by simp - next - case TRUE - with l u `b1\vbits` show ?thesis using assms by simp - next - case FALSE - with l u `b1\vbits` show ?thesis using assms by simp - next - case (LVAL x7) - with l u `b1\vbits` show ?thesis using assms by simp - next - case (PLUS x81 x82) - with l u `b1\vbits` show ?thesis using assms by simp - next - case (MINUS x91 x92) - with l u `b1\vbits` show ?thesis using assms by simp - next - case (EQUAL x101 x102) - with l u `b1\vbits` show ?thesis using assms by simp - next - case (LESS x111 x112) - with l u `b1\vbits` show ?thesis using assms by simp - next - case (AND x121 x122) - with l u `b1\vbits` show ?thesis using assms by simp - next - case (OR x131 x132) - with l u `b1\vbits` show ?thesis using assms by simp - next - case (NOT x131) - with l u `b1\vbits` show ?thesis using assms by simp - next - case (CALL x181 x182) - with l u `b1\vbits` show ?thesis using assms by simp - next - case (ECALL x191 x192 x193 x194) - with l u `b1\vbits` show ?thesis using assms by simp - qed - next - assume "\ b1\vbits" - with l u show ?thesis using assms by simp - qed - next - case (ADDRESS x3) - with l show ?thesis using assms by simp - next - case (BALANCE x4) - with l show ?thesis using assms by simp - next - case THIS - with l show ?thesis using assms by simp - next - case SENDER - with l show ?thesis using assms by simp - next - case VALUE - with l show ?thesis using assms by simp - next - case TRUE - with l show ?thesis using assms by simp - next - case FALSE - with l show ?thesis using assms by simp - next - case (LVAL x7) - with l show ?thesis using assms by simp - next - case (PLUS x81 x82) - with l show ?thesis using assms by simp - next - case (MINUS x91 x92) - with l show ?thesis using assms by simp - next - case (EQUAL x101 x102) - with l show ?thesis using assms by simp - next - case (LESS x111 x112) - with l show ?thesis using assms by simp - next - case (AND x121 x122) - with l show ?thesis using assms by simp - next - case (OR x131 x132) - with l show ?thesis using assms by simp - next - case (NOT x131) - with l show ?thesis using assms by simp - next - case (CALL x181 x182) - with l show ?thesis using assms by simp - next - case (ECALL x191 x192 x193 x194) - with l show ?thesis using assms by simp - qed -next - case a: (AND e1 e2) - show ?thesis - proof (cases "eupdate e1") - case (INT x11 x12) - with a show ?thesis using assms by simp - next - case (UINT x21 x22) - with a show ?thesis using assms by simp - next - case (ADDRESS x3) - with a show ?thesis using assms by simp - next - case (BALANCE x4) - with a show ?thesis using assms by simp - next - case THIS - with a show ?thesis using assms by simp - next - case SENDER - with a show ?thesis using assms by simp - next - case VALUE - with a show ?thesis using assms by simp - next - case t: TRUE - show ?thesis - proof (cases "eupdate e2") - case (INT x11 x12) - with a t show ?thesis using assms by simp - next - case (UINT x21 x22) - with a t show ?thesis using assms by simp - next - case (ADDRESS x3) - with a t show ?thesis using assms by simp - next - case (BALANCE x4) - with a t show ?thesis using assms by simp - next - case THIS - with a t show ?thesis using assms by simp - next - case SENDER - with a t show ?thesis using assms by simp - next - case VALUE - with a t show ?thesis using assms by simp - next - case TRUE - with a t show ?thesis using assms by simp - next - case FALSE - with a t show ?thesis using assms by simp - next - case (LVAL x7) - with a t show ?thesis using assms by simp - next - case (PLUS x81 x82) - with a t show ?thesis using assms by simp - next - case (MINUS x91 x92) - with a t show ?thesis using assms by simp - next - case (EQUAL x101 x102) - with a t show ?thesis using assms by simp - next - case (LESS x111 x112) - with a t show ?thesis using assms by simp - next - case (AND x121 x122) - with a t show ?thesis using assms by simp - next - case (OR x131 x132) - with a t show ?thesis using assms by simp - next - case (NOT x131) - with a t show ?thesis using assms by simp - next - case (CALL x181 x182) - with a t show ?thesis using assms by simp - next - case (ECALL x191 x192 x193 x194) - with a t show ?thesis using assms by simp - qed - next - case f: FALSE - show ?thesis - proof (cases "eupdate e2") - case (INT x11 x12) - with a f show ?thesis using assms by simp - next - case (UINT x21 x22) - with a f show ?thesis using assms by simp - next - case (ADDRESS x3) - with a f show ?thesis using assms by simp - next - case (BALANCE x4) - with a f show ?thesis using assms by simp - next - case THIS - with a f show ?thesis using assms by simp - next - case SENDER - with a f show ?thesis using assms by simp - next - case VALUE - with a f show ?thesis using assms by simp - next - case TRUE - with a f show ?thesis using assms by simp - next - case FALSE - with a f show ?thesis using assms by simp - next - case (LVAL x7) - with a f show ?thesis using assms by simp - next - case (PLUS x81 x82) - with a f show ?thesis using assms by simp - next - case (MINUS x91 x92) - with a f show ?thesis using assms by simp - next - case (EQUAL x101 x102) - with a f show ?thesis using assms by simp - next - case (LESS x111 x112) - with a f show ?thesis using assms by simp - next - case (AND x121 x122) - with a f show ?thesis using assms by simp - next - case (OR x131 x132) - with a f show ?thesis using assms by simp - next - case (NOT x131) - with a f show ?thesis using assms by simp - next - case (CALL x181 x182) - with a f show ?thesis using assms by simp - next - case (ECALL x191 x192 x193 x194) - with a f show ?thesis using assms by simp - qed - next - case (LVAL x7) - with a show ?thesis using assms by simp - next - case (PLUS x81 x82) - with a show ?thesis using assms by simp - next - case (MINUS x91 x92) - with a show ?thesis using assms by simp - next - case (EQUAL x101 x102) - with a show ?thesis using assms by simp - next - case (LESS x111 x112) - with a show ?thesis using assms by simp - next - case (AND x121 x122) - with a show ?thesis using assms by simp - next - case (OR x131 x132) - with a show ?thesis using assms by simp - next - case (NOT x131) - with a show ?thesis using assms by simp - next - case (CALL x181 x182) - with a show ?thesis using assms by simp - next - case (ECALL x191 x192 x193 x194) - with a show ?thesis using assms by simp - qed -next - case o: (OR e1 e2) - show ?thesis - proof (cases "eupdate e1") - case (INT x11 x12) - with o show ?thesis using assms by simp - next - case (UINT x21 x22) - with o show ?thesis using assms by simp - next - case (ADDRESS x3) - with o show ?thesis using assms by simp - next - case (BALANCE x4) - with o show ?thesis using assms by simp - next - case THIS - with o show ?thesis using assms by simp - next - case SENDER - with o show ?thesis using assms by simp - next - case VALUE - with o show ?thesis using assms by simp - next - case t: TRUE - show ?thesis - proof (cases "eupdate e2") - case (INT x11 x12) - with o t show ?thesis using assms by simp - next - case (UINT x21 x22) - with o t show ?thesis using assms by simp - next - case (ADDRESS x3) - with o t show ?thesis using assms by simp - next - case (BALANCE x4) - with o t show ?thesis using assms by simp - next - case THIS - with o t show ?thesis using assms by simp - next - case SENDER - with o t show ?thesis using assms by simp - next - case VALUE - with o t show ?thesis using assms by simp - next - case TRUE - with o t show ?thesis using assms by simp - next - case FALSE - with o t show ?thesis using assms by simp - next - case (LVAL x7) - with o t show ?thesis using assms by simp - next - case (PLUS x81 x82) - with o t show ?thesis using assms by simp - next - case (MINUS x91 x92) - with o t show ?thesis using assms by simp - next - case (EQUAL x101 x102) - with o t show ?thesis using assms by simp - next - case (LESS x111 x112) - with o t show ?thesis using assms by simp - next - case (AND x121 x122) - with o t show ?thesis using assms by simp - next - case (OR x131 x132) - with o t show ?thesis using assms by simp - next - case (NOT x131) - with o t show ?thesis using assms by simp - next - case (CALL x181 x182) - with o t show ?thesis using assms by simp - next - case (ECALL x191 x192 x193 x194) - with o t show ?thesis using assms by simp - qed - next - case f: FALSE - show ?thesis - proof (cases "eupdate e2") - case (INT x11 x12) - with o f show ?thesis using assms by simp - next - case (UINT x21 x22) - with o f show ?thesis using assms by simp - next - case (ADDRESS x3) - with o f show ?thesis using assms by simp - next - case (BALANCE x4) - with o f show ?thesis using assms by simp - next - case THIS - with o f show ?thesis using assms by simp - next - case SENDER - with o f show ?thesis using assms by simp - next - case VALUE - with o f show ?thesis using assms by simp - next - case TRUE - with o f show ?thesis using assms by simp - next - case FALSE - with o f show ?thesis using assms by simp - next - case (LVAL x7) - with o f show ?thesis using assms by simp - next - case (PLUS x81 x82) - with o f show ?thesis using assms by simp - next - case (MINUS x91 x92) - with o f show ?thesis using assms by simp - next - case (EQUAL x101 x102) - with o f show ?thesis using assms by simp - next - case (LESS x111 x112) - with o f show ?thesis using assms by simp - next - case (AND x121 x122) - with o f show ?thesis using assms by simp - next - case (OR x131 x132) - with o f show ?thesis using assms by simp - next - case (NOT x131) - with o f show ?thesis using assms by simp - next - case (CALL x181 x182) - with o f show ?thesis using assms by simp - next - case (ECALL x191 x192 x193 x194) - with o f show ?thesis using assms by simp - qed - next - case (LVAL x7) - with o show ?thesis using assms by simp - next - case (PLUS x81 x82) - with o show ?thesis using assms by simp - next - case (MINUS x91 x92) - with o show ?thesis using assms by simp - next - case (EQUAL x101 x102) - with o show ?thesis using assms by simp - next - case (LESS x111 x112) - with o show ?thesis using assms by simp - next - case (AND x121 x122) - with o show ?thesis using assms by simp - next - case (OR x131 x132) - with o show ?thesis using assms by simp - next - case (NOT x131) - with o show ?thesis using assms by simp - next - case (CALL x181 x182) - with o show ?thesis using assms by simp - next - case (ECALL x191 x192 x193 x194) - with o show ?thesis using assms by simp - qed -next - case o: (NOT x) - show ?thesis - proof (cases "eupdate x") - case (INT x11 x12) - with o show ?thesis using assms by simp - next - case (UINT x21 x22) - with o show ?thesis using assms by simp - next - case (ADDRESS x3) - with o show ?thesis using assms by simp - next - case (BALANCE x4) - with o show ?thesis using assms by simp - next - case THIS - with o show ?thesis using assms by simp - next - case SENDER - with o show ?thesis using assms by simp - next - case VALUE - with o show ?thesis using assms by simp - next - case t: TRUE - with o show ?thesis using assms by simp - next - case f: FALSE - with o show ?thesis using assms by simp - next - case (LVAL x7) - with o show ?thesis using assms by simp - next - case (PLUS x81 x82) - with o show ?thesis using assms by simp - next - case (MINUS x91 x92) - with o show ?thesis using assms by simp - next - case (EQUAL x101 x102) - with o show ?thesis using assms by simp - next - case (LESS x111 x112) - with o show ?thesis using assms by simp - next - case (AND x121 x122) - with o show ?thesis using assms by simp - next - case (OR x131 x132) - with o show ?thesis using assms by simp - next - case (NOT x131) - with o show ?thesis using assms by simp - next - case (CALL x181 x182) - with o show ?thesis using assms by simp - next - case (ECALL x191 x192 x193 x194) - with o show ?thesis using assms by simp - qed -next - case (CALL x181 x182) - with assms show ?thesis by simp -next - case (ECALL x191 x192 x193 x194) - with assms show ?thesis by simp -qed - -lemma no_gas: - assumes "\ gas st > 0" - shows "expr ex ep env cd st = Exception Gas" -proof (cases ex) - case (INT x11 x12) - with assms show ?thesis by simp -next - case (UINT x21 x22) - with assms show ?thesis by simp -next - case (ADDRESS x3) - with assms show ?thesis by simp -next - case (BALANCE x4) - with assms show ?thesis by simp -next - case THIS - with assms show ?thesis by simp -next - case SENDER - with assms show ?thesis by simp -next - case VALUE - with assms show ?thesis by simp -next - case TRUE - with assms show ?thesis by simp -next - case FALSE - with assms show ?thesis by simp -next - case (LVAL x10) - with assms show ?thesis by simp -next - case (PLUS x111 x112) - with assms show ?thesis by simp -next - case (MINUS x121 x122) - with assms show ?thesis by simp -next - case (EQUAL x131 x132) - with assms show ?thesis by simp -next - case (LESS x141 x142) - with assms show ?thesis by simp -next - case (AND x151 x152) - with assms show ?thesis by simp -next - case (OR x161 x162) - with assms show ?thesis by simp -next - case (NOT x17) - with assms show ?thesis by simp -next - case (CALL x181 x182) - with assms show ?thesis by simp -next - case (ECALL x191 x192 x193 x194) - with assms show ?thesis by simp -qed - -lemma lift_eq: - assumes "expr e1 ep env cd st = expr e1' ep env cd st" - and "\st' rv. expr e1 ep env cd st = Normal (rv, st') \ expr e2 ep env cd st'= expr e2' ep env cd st'" - shows "lift expr f e1 e2 ep env cd st=lift expr f e1' e2' ep env cd st" -proof (cases "expr e1 ep env cd st") - case s1: (n a st') - then show ?thesis - proof (cases a) - case f1:(Pair a b) - then show ?thesis - proof (cases a) - case k1:(KValue x1) - then show ?thesis - proof (cases b) - case v1: (Value x1) - then show ?thesis - proof (cases "expr e2 ep env cd st'") - case s2: (n a' st'') - then show ?thesis - proof (cases a') - case f2:(Pair a' b') - then show ?thesis - proof (cases a') - case (KValue x1') - with s1 f1 k1 v1 assms(1) assms(2) show ?thesis by auto - next - case (KCDptr x2) - with s1 f1 k1 v1 assms(1) assms(2) show ?thesis by auto - next - case (KMemptr x2') - with s1 f1 k1 v1 assms(1) assms(2) show ?thesis by auto - next - case (KStoptr x3') - with s1 f1 k1 v1 assms(1) assms(2) show ?thesis by auto - qed - qed - next - case (e e) - then show ?thesis using k1 s1 v1 assms(1) assms(2) f1 by auto - qed - next - case (Calldata x2) - then show ?thesis using k1 s1 assms(1) f1 by auto - next - case (Memory x2) - then show ?thesis using k1 s1 assms(1) f1 by auto - next - case (Storage x3) - then show ?thesis using k1 s1 assms(1) f1 by auto - qed - next - case (KCDptr x2) - then show ?thesis using s1 assms(1) f1 by fastforce - next - case (KMemptr x2) - then show ?thesis using s1 assms(1) f1 by fastforce - next - case (KStoptr x3) - then show ?thesis using s1 assms(1) f1 by fastforce - qed - qed -next - case (e e) - then show ?thesis using assms(1) by simp -qed - -lemma ssel_eq_ssel: - "(\i st. i \ set ix \ expr i ep env cd st = expr (f i) ep env cd st) - \ ssel tp loc ix ep env cd st = ssel tp loc (map f ix) ep env cd st" -proof (induction ix arbitrary: tp loc ep env cd st) - case Nil - then show ?case by simp -next - case c1: (Cons i ix) - then show ?case - proof (cases tp) - case tp1: (STArray al tp) - then show ?thesis - proof (cases "expr i ep env cd st") - case s1: (n a st') - then show ?thesis - proof (cases a) - case f1: (Pair a b) - then show ?thesis - proof (cases a) - case k1: (KValue v) - then show ?thesis - proof (cases b) - case v1: (Value t) - then show ?thesis - proof (cases "less t (TUInt 256) v (ShowL\<^sub>i\<^sub>n\<^sub>t al)") - case None - with v1 k1 tp1 s1 c1.prems f1 show ?thesis by simp - next - case s2: (Some a) - then show ?thesis - proof (cases a) - case p1: (Pair a b) - then show ?thesis - proof (cases b) - case (TSInt x1) - with s2 p1 v1 k1 tp1 s1 c1.prems f1 show ?thesis by simp - next - case (TUInt x2) - with s2 p1 v1 k1 tp1 s1 c1.prems f1 show ?thesis by simp - next - case b1: TBool - show ?thesis - proof cases - assume "a = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True" - from c1.IH[OF c1.prems] have - "ssel tp (hash loc v) ix ep env cd st' = ssel tp (hash loc v) (map f ix) ep env cd st'" - by simp - with mp s2 b1 p1 v1 k1 tp1 s1 c1.prems f1 show ?thesis by simp - next - assume "\ a = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True" - with s2 p1 v1 k1 tp1 s1 c1.prems f1 show ?thesis by simp - qed - next - case TAddr - with s2 p1 v1 k1 tp1 s1 c1.prems f1 show ?thesis by simp - qed - qed - qed - next - case (Calldata x2) - with k1 tp1 s1 c1.prems f1 show ?thesis by simp - next - case (Memory x2) - with k1 tp1 s1 c1.prems f1 show ?thesis by simp - next - case (Storage x3) - with k1 tp1 s1 c1.prems f1 show ?thesis by simp - qed - next - case (KCDptr x2) - with tp1 s1 c1.prems f1 show ?thesis by simp - next - case (KMemptr x2) - with tp1 s1 c1.prems f1 show ?thesis by simp - next - case (KStoptr x3) - with tp1 s1 c1.prems f1 show ?thesis by simp - qed - qed - next - case (e e) - with tp1 c1.prems show ?thesis by simp - qed - next - case tp1: (STMap _ t) - then show ?thesis - proof (cases "expr i ep env cd st") - case s1: (n a s) - then show ?thesis - proof (cases a) - case f1: (Pair a b) - then show ?thesis - proof (cases a) - case k1: (KValue v) - from c1.IH[OF c1.prems] have - "ssel tp (hash loc v) ix ep env cd st = ssel tp (hash loc v) (map f ix) ep env cd st" by simp - with k1 tp1 s1 c1 f1 show ?thesis by simp - next - case (KCDptr x2) - with tp1 s1 c1.prems f1 show ?thesis by simp - next - case (KMemptr x2) - with tp1 s1 c1.prems f1 show ?thesis by simp - next - case (KStoptr x3) - with tp1 s1 c1.prems f1 show ?thesis by simp - qed - qed - next - case (e e) - with tp1 c1.prems show ?thesis by simp - qed - next - case (STValue x2) - then show ?thesis by simp - qed -qed - -lemma msel_eq_msel: -"(\i st. i \ set ix \ expr i ep env cd st = expr (f i) ep env cd st) \ - msel c tp loc ix ep env cd st = msel c tp loc (map f ix) ep env cd st" -proof (induction ix arbitrary: c tp loc ep env cd st) - case Nil - then show ?case by simp -next - case c1: (Cons i ix) - then show ?case - proof (cases tp) - case tp1: (MTArray al tp) - then show ?thesis - proof (cases ix) - case Nil - thus ?thesis using tp1 c1.prems by auto - next - case c2: (Cons a list) - then show ?thesis - proof (cases "expr i ep env cd st") - case s1: (n a st') - then show ?thesis - proof (cases a) - case f1: (Pair a b) - then show ?thesis - proof (cases a) - case k1: (KValue v) - then show ?thesis - proof (cases b) - case v1: (Value t) - then show ?thesis - proof (cases "less t (TUInt 256) v (ShowL\<^sub>i\<^sub>n\<^sub>t al)") - case None - with v1 k1 tp1 s1 c1.prems f1 show ?thesis using c2 by simp - next - case s2: (Some a) - then show ?thesis - proof (cases a) - case p1: (Pair a b) - then show ?thesis - proof (cases b) - case (TSInt x1) - with s2 p1 v1 k1 tp1 s1 c1.prems f1 show ?thesis using c2 by simp - next - case (TUInt x2) - with s2 p1 v1 k1 tp1 s1 c1.prems f1 show ?thesis using c2 by simp - next - case b1: TBool - show ?thesis - proof cases - assume "a = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True" - then show ?thesis - proof (cases c) - case True - then show ?thesis - proof (cases "accessStore (hash loc v) (memory st')") - case None - with s2 b1 p1 v1 k1 tp1 s1 c1.prems f1 True show ?thesis using c2 by simp - next - case s3: (Some a) - then show ?thesis - proof (cases a) - case (MValue x1) - with s2 s3 b1 p1 v1 k1 tp1 s1 c1.prems f1 True show ?thesis using c2 by simp - next - case mp: (MPointer l) - from c1.IH[OF c1.prems] - have "msel c tp l ix ep env cd st' = msel c tp l (map f ix) ep env cd st'" by simp - with mp s2 s3 b1 p1 v1 k1 tp1 s1 c1.prems f1 True show ?thesis using c2 by simp - qed - qed - next - case False - then show ?thesis - proof (cases "accessStore (hash loc v) cd") - case None - with s2 b1 p1 v1 k1 tp1 s1 c1.prems f1 False show ?thesis using c2 by simp - next - case s3: (Some a) - then show ?thesis - proof (cases a) - case (MValue x1) - with s2 s3 b1 p1 v1 k1 tp1 s1 c1.prems f1 False show ?thesis using c2 by simp - next - case mp: (MPointer l) - from c1.IH[OF c1.prems] - have "msel c tp l ix ep env cd st' = msel c tp l (map f ix) ep env cd st'" by simp - with mp s2 s3 b1 p1 v1 k1 tp1 s1 c1.prems f1 False show ?thesis using c2 by simp - qed - qed - qed - next - assume "\ a = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True" - with s2 p1 v1 k1 tp1 s1 c1.prems f1 show ?thesis using c2 by simp - qed - next - case TAddr - with s2 p1 v1 k1 tp1 s1 c1.prems f1 show ?thesis using c2 by simp - qed - qed - qed - next - case (Calldata x2) - with k1 tp1 s1 c1.prems f1 show ?thesis using c2 by simp - next - case (Memory x2) - with k1 tp1 s1 c1.prems f1 show ?thesis using c2 by simp - next - case (Storage x3) - with k1 tp1 s1 c1.prems f1 show ?thesis using c2 by simp - qed - next - case (KCDptr x2) - with tp1 s1 c1.prems f1 show ?thesis using c2 by simp - next - case (KMemptr x2) - with tp1 s1 c1.prems f1 show ?thesis using c2 by simp - next - case (KStoptr x3) - with tp1 s1 c1.prems f1 show ?thesis using c2 by simp - qed - qed - next - case (e e) - with tp1 c1.prems show ?thesis using c2 by simp - qed - qed - next - case (MTValue x2) - then show ?thesis by simp - qed -qed - -lemma ref_eq: - assumes "\e st. e \ set ex \ expr e ep env cd st = expr (f e) ep env cd st" - shows "rexp (Ref i ex) ep env cd st=rexp (Ref i (map f ex)) ep env cd st" -proof (cases "fmlookup (denvalue env) i") - case None - then show ?thesis by simp -next - case s1: (Some a) - then show ?thesis - proof (cases a) - case p1: (Pair tp b) - then show ?thesis - proof (cases b) - case k1: (Stackloc l) - then show ?thesis - proof (cases "accessStore l (stack st)") - case None - with s1 p1 k1 show ?thesis by simp - next - case s2: (Some a') - then show ?thesis - proof (cases a') - case (KValue _) - with s1 s2 p1 k1 show ?thesis by simp - next - case cp: (KCDptr cp) - then show ?thesis - proof (cases tp) - case (Value x1) - with mp s1 s2 p1 k1 show ?thesis by simp - next - case mt: (Calldata ct) - from msel_eq_msel have - "msel False ct cp ex ep env cd st=msel False ct cp (map f ex) ep env cd st" using assms by blast - thus ?thesis using s1 s2 p1 k1 mt cp by simp - next - case mt: (Memory mt) - from msel_eq_msel have - "msel True mt cp ex ep env cd st=msel True mt cp (map f ex) ep env cd st" using assms by blast - thus ?thesis using s1 s2 p1 k1 mt cp by simp - next - case (Storage x3) - with cp s1 s2 p1 k1 show ?thesis by simp - qed - next - case mp: (KMemptr mp) - then show ?thesis - proof (cases tp) - case (Value x1) - with mp s1 s2 p1 k1 show ?thesis by simp - next - case mt: (Calldata ct) - from msel_eq_msel have - "msel True ct mp ex ep env cd st=msel True ct mp (map f ex) ep env cd st" using assms by blast - thus ?thesis using s1 s2 p1 k1 mt mp by simp - next - case mt: (Memory mt) - from msel_eq_msel have - "msel True mt mp ex ep env cd st=msel True mt mp (map f ex) ep env cd st" using assms by blast - thus ?thesis using s1 s2 p1 k1 mt mp by simp - next - case (Storage x3) - with mp s1 s2 p1 k1 show ?thesis by simp - qed - next - case sp: (KStoptr sp) - then show ?thesis - proof (cases tp) - case (Value x1) - then show ?thesis using s1 s2 p1 k1 sp by simp - next - case (Calldata x2) - then show ?thesis using s1 s2 p1 k1 sp by simp - next - case (Memory x2) - then show ?thesis using s1 s2 p1 k1 sp by simp - next - case st: (Storage stp) - from ssel_eq_ssel have - "ssel stp sp ex ep env cd st=ssel stp sp (map f ex) ep env cd st" using assms by blast - thus ?thesis using s1 s2 p1 k1 st sp by simp - qed - qed - qed - next - case sl:(Storeloc sl) - then show ?thesis - proof (cases tp) - case (Value x1) - then show ?thesis using s1 p1 sl by simp - next - case (Calldata x2) - then show ?thesis using s1 p1 sl by simp - next - case (Memory x2) - then show ?thesis using s1 p1 sl by simp - next - case st: (Storage stp) - from ssel_eq_ssel have - "ssel stp sl ex ep env cd st=ssel stp sl (map f ex) ep env cd st" using assms by blast - thus ?thesis using s1 sl p1 st by simp - qed - qed - qed -qed - -text\ - The following theorem proves that the update function preserves the semantics of expressions. -\ -theorem update_correctness: - "\st lb lv. expr ex ep env cd st = expr (eupdate ex) ep env cd st" - "\st. rexp lv ep env cd st = rexp (lupdate lv) ep env cd st" -proof (induction ex and lv) - case (Id x) - then show ?case by simp -next - case (Ref d ix) - then show ?case using ref_eq[where f="eupdate"] by simp -next - case (INT b v) - then show ?case - proof (cases "gas st > 0") - case True - then show ?thesis - proof cases - assume "b\vbits" - show ?thesis - proof cases - let ?m_def = "(-(2^(b-1)) + (v+2^(b-1)) mod (2^b))" - assume "v \ 0" - - from `b\vbits` True have - "expr (E.INT b v) ep env cd st = Normal ((KValue (createSInt b v), Value (TSInt b)), st)" by simp - also have "createSInt b v = createSInt b ?m_def" using `b\vbits` `v \ 0` by auto - also from `v \ 0` `b\vbits` True have - "Normal ((KValue (createSInt b ?m_def), Value (TSInt b)),st) = expr (eupdate (E.INT b v)) ep env cd st" - by simp - finally show "expr (E.INT b v) ep env cd st = expr (eupdate (E.INT b v)) ep env cd st" by simp - next - let ?m_def = "(2^(b-1) - (-v+2^(b-1)-1) mod (2^b) - 1)" - assume "\ v \ 0" - - from `b\vbits` True have - "expr (E.INT b v) ep env cd st = Normal ((KValue (createSInt b v), Value (TSInt b)), st)" by simp - also have "createSInt b v = createSInt b ?m_def" using `b\vbits` `\ v \ 0` by auto - also from `\ v \ 0` `b\vbits` True have - "Normal ((KValue (createSInt b ?m_def), Value (TSInt b)),st) =expr (eupdate (E.INT b v)) ep env cd st" - by simp - finally show "expr (E.INT b v) ep env cd st = expr (eupdate (E.INT b v)) ep env cd st" by simp - qed - next - assume "\ b\vbits" - thus ?thesis by auto - qed - next - case False - then show ?thesis using no_gas by simp - qed -next - case (UINT x1 x2) - then show ?case by simp -next - case (ADDRESS x) - then show ?case by simp -next - case (BALANCE x) - then show ?case by simp -next - case THIS - then show ?case by simp -next - case SENDER - then show ?case by simp -next - case VALUE - then show ?case by simp -next - case TRUE - then show ?case by simp -next - case FALSE - then show ?case by simp -next - case (LVAL x) - then show ?case by simp -next - case p: (PLUS e1 e2) - show ?case - proof (cases "eupdate e1") - case i: (INT b1 v1) - with p.IH have expr1: "expr e1 ep env cd st = expr (E.INT b1 v1) ep env cd st" by simp - then show ?thesis - proof (cases "gas st > 0") - case True - then show ?thesis - proof (cases) - assume "b1 \ vbits" - with expr1 True - have "expr e1 ep env cd st=Normal ((KValue (createSInt b1 v1), Value (TSInt b1)),st)" by simp - moreover from i `b1 \ vbits` - have "v1 < 2^(b1-1)" and "v1 \ -(2^(b1-1))" using update_bounds_int by auto - moreover from `b1 \ vbits` have "0 < b1" by auto - ultimately have r1: "expr e1 ep env cd st = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t v1), Value (TSInt b1)),st)" - using createSInt_id[of v1 b1] by simp - thus ?thesis - proof (cases "eupdate e2") - case i2: (INT b2 v2) - with p.IH have expr2: "expr e2 ep env cd st = expr (E.INT b2 v2) ep env cd st" by simp - then show ?thesis - proof (cases) - let ?v="v1 + v2" - assume "b2 \ vbits" - with expr2 True - have "expr e2 ep env cd st=Normal ((KValue (createSInt b2 v2), Value (TSInt b2)),st)" by simp - moreover from i2 `b2 \ vbits` - have "v2 < 2^(b2-1)" and "v2 \ -(2^(b2-1))" using update_bounds_int by auto - moreover from `b2 \ vbits` have "0 < b2" by auto - ultimately have r2: "expr e2 ep env cd st = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t v2), Value (TSInt b2)),st)" - using createSInt_id[of v2 b2] by simp - thus ?thesis - proof (cases) - let ?x="- (2 ^ (max b1 b2 - 1)) + (?v + 2 ^ (max b1 b2 - 1)) mod 2 ^ max b1 b2" - assume "?v\0" - hence "createSInt (max b1 b2) ?v = (ShowL\<^sub>i\<^sub>n\<^sub>t ?x)" by simp - moreover have "add (TSInt b1) (TSInt b2) (ShowL\<^sub>i\<^sub>n\<^sub>t v1) (ShowL\<^sub>i\<^sub>n\<^sub>t v2) - = Some (createSInt (max b1 b2) ?v, TSInt (max b1 b2))" - using Read_ShowL_id add_def olift.simps(1)[of "(+)" b1 b2] by simp - ultimately have "expr (PLUS e1 e2) ep env cd st - = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TSInt (max b1 b2))),st)" using r1 r2 True by simp - moreover have "expr (eupdate (PLUS e1 e2)) ep env cd st - = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TSInt (max b1 b2))),st)" - proof - - from `b1 \ vbits` `b2 \ vbits` `?v\0` - have "eupdate (PLUS e1 e2) = E.INT (max b1 b2) ?x" using i i2 by simp - moreover have "expr (E.INT (max b1 b2) ?x) ep env cd st - = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TSInt (max b1 b2))),st)" - proof - - from `b1 \ vbits` `b2 \ vbits` have "max b1 b2 \ vbits" using vbits_max by simp - with True have "expr (E.INT (max b1 b2) ?x) ep env cd st - = Normal ((KValue (createSInt (max b1 b2) ?x), Value (TSInt (max b1 b2))),st)" by simp - moreover from `0 < b1` - have "?x < 2 ^ (max b1 b2 - 1)" using upper_bound3 by simp - moreover from `0 < b1` have "0 < max b1 b2" using max_def by simp - ultimately show ?thesis using createSInt_id[of ?x "max b1 b2"] by simp - qed - ultimately show ?thesis by simp - qed - ultimately show ?thesis by simp - next - let ?x="2^(max b1 b2 -1) - (-?v+2^(max b1 b2-1)-1) mod (2^max b1 b2) - 1" - assume "\ ?v\0" - hence "createSInt (max b1 b2) ?v = (ShowL\<^sub>i\<^sub>n\<^sub>t ?x)" by simp - moreover have "add (TSInt b1) (TSInt b2) (ShowL\<^sub>i\<^sub>n\<^sub>t v1) (ShowL\<^sub>i\<^sub>n\<^sub>t v2) - = Some (createSInt (max b1 b2) ?v, TSInt (max b1 b2))" - using Read_ShowL_id add_def olift.simps(1)[of "(+)" b1 b2] by simp - ultimately have "expr (PLUS e1 e2) ep env cd st - = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TSInt (max b1 b2))),st)" using True r1 r2 by simp - moreover have "expr (eupdate (PLUS e1 e2)) ep env cd st - = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TSInt (max b1 b2))),st)" - proof - - from `b1 \ vbits` `b2 \ vbits` `\?v\0` - have "eupdate (PLUS e1 e2) = E.INT (max b1 b2) ?x" using i i2 by simp - moreover have "expr (E.INT (max b1 b2) ?x) ep env cd st - = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TSInt (max b1 b2))),st)" - proof - - from `b1 \ vbits` `b2 \ vbits` have "max b1 b2 \ vbits" using vbits_max by simp - with True have "expr (E.INT (max b1 b2) ?x) ep env cd st - = Normal ((KValue (createSInt (max b1 b2) ?x), Value (TSInt (max b1 b2))),st)" by simp - moreover from `0 < b1` - have "?x \ - (2 ^ (max b1 b2 - 1))" using lower_bound2[of "max b1 b2" ?v] by simp - moreover from `b1 > 0` have "2^(max b1 b2 -1) > (0::nat)" by simp - hence "2^(max b1 b2 -1) - (-?v+2^(max b1 b2-1)-1) mod (2^max b1 b2) - 1 < 2 ^ (max b1 b2 - 1)" - by (simp add: algebra_simps flip: zle_diff1_eq) - moreover from `0 < b1` have "0 < max b1 b2" using max_def by simp - ultimately show ?thesis using createSInt_id[of ?x "max b1 b2"] by simp - qed - ultimately show ?thesis by simp - qed - ultimately show ?thesis by simp - qed - next - assume "\ b2 \ vbits" - with p i i2 show ?thesis by simp - qed - next - case u2: (UINT b2 v2) - with p.IH have expr2: "expr e2 ep env cd st = expr (UINT b2 v2) ep env cd st" by simp - then show ?thesis - proof (cases) - let ?v="v1 + v2" - assume "b2 \ vbits" - with expr2 True - have "expr e2 ep env cd st=Normal ((KValue (createUInt b2 v2), Value (TUInt b2)),st)" by simp - moreover from u2 `b2 \ vbits` - have "v2 < 2^b2" and "v2 \ 0" using update_bounds_uint by auto - moreover from `b2 \ vbits` have "0 < b2" by auto - ultimately have r2: "expr e2 ep env cd st = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t v2), Value (TUInt b2)),st)" - using createUInt_id[of v2 b2] by simp - thus ?thesis - proof (cases) - assume "b20" - hence "createSInt b1 ?v = (ShowL\<^sub>i\<^sub>n\<^sub>t ?x)" using `b2i\<^sub>n\<^sub>t v1) (ShowL\<^sub>i\<^sub>n\<^sub>t v2) - = Some (createSInt b1 ?v, TSInt b1)" - using Read_ShowL_id add_def olift.simps(3)[of "(+)" b1 b2] `b2i\<^sub>n\<^sub>t ?x), Value (TSInt b1)),st)" using r1 r2 True by simp - moreover have "expr (eupdate (PLUS e1 e2)) ep env cd st - = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TSInt b1)),st)" - proof - - from `b1 \ vbits` `b2 \ vbits` `?v\0` `b2i\<^sub>n\<^sub>t ?x), Value (TSInt b1)),st)" - proof - - from `b1 \ vbits` True have "expr (E.INT b1 ?x) ep env cd st - = Normal ((KValue (createSInt b1 ?x), Value (TSInt b1)),st)" by simp - moreover from `0 < b1` have "?x < 2 ^ (b1 - 1)" using upper_bound2 by simp - ultimately show ?thesis using createSInt_id[of ?x "b1"] `0 < b1` by simp - qed - ultimately show ?thesis by simp - qed - ultimately show ?thesis by simp - next - let ?x="2^(b1 -1) - (-?v+2^(b1-1)-1) mod (2^b1) - 1" - assume "\ ?v\0" - hence "createSInt b1 ?v = (ShowL\<^sub>i\<^sub>n\<^sub>t ?x)" by simp - moreover have "add (TSInt b1) (TUInt b2) (ShowL\<^sub>i\<^sub>n\<^sub>t v1) (ShowL\<^sub>i\<^sub>n\<^sub>t v2) - = Some (createSInt b1 ?v, TSInt b1)" - using Read_ShowL_id add_def olift.simps(3)[of "(+)" b1 b2] `b2i\<^sub>n\<^sub>t ?x), Value (TSInt b1)),st)" using r1 r2 True by simp - moreover have "expr (eupdate (PLUS e1 e2)) ep env cd st - = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TSInt b1)),st)" - proof - - from `b1 \ vbits` `b2 \ vbits` `\?v\0` `b2i\<^sub>n\<^sub>t ?x), Value (TSInt b1)),st)" - proof - - from `b1 \ vbits` True have "expr (E.INT b1 ?x) ep env cd st - = Normal ((KValue (createSInt b1 ?x), Value (TSInt b1)),st)" by simp - moreover from `0 < b1` have "?x \ - (2 ^ (b1 - 1))" using upper_bound2 by simp - moreover have "2^(b1-1) - (-?v+2^(b1-1)-1) mod (2^b1) - 1 < 2 ^ (b1 - 1)" - by (simp add: algebra_simps flip: int_one_le_iff_zero_less) - ultimately show ?thesis using createSInt_id[of ?x b1] `0 < b1` by simp - qed - ultimately show ?thesis by simp - qed - ultimately show ?thesis by simp - qed - next - assume "\ b2 < b1" - with p i u2 show ?thesis by simp - qed - next - assume "\ b2 \ vbits" - with p i u2 show ?thesis by simp - qed - next - case (ADDRESS _) - with p i show ?thesis by simp - next - case (BALANCE _) - with p i show ?thesis by simp - next - case THIS - with p i show ?thesis by simp - next - case SENDER - with p i show ?thesis by simp - next - case VALUE - with p i show ?thesis by simp - next - case TRUE - with p i show ?thesis by simp - next - case FALSE - with p i show ?thesis by simp - next - case (LVAL _) - with p i show ?thesis by simp - next - case (PLUS _ _) - with p i show ?thesis by simp - next - case (MINUS _ _) - with p i show ?thesis by simp - next - case (EQUAL _ _) - with p i show ?thesis by simp - next - case (LESS _ _) - with p i show ?thesis by simp - next - case (AND _ _) - with p i show ?thesis by simp - next - case (OR _ _) - with p i show ?thesis by simp - next - case (NOT _) - with p i show ?thesis by simp - next - case (CALL x181 x182) - with p i show ?thesis by simp - next - case (ECALL x191 x192 x193 x194) - with p i show ?thesis by simp - qed - next - assume "\ b1 \ vbits" - with p i show ?thesis by simp - qed - next - case False - then show ?thesis using no_gas by simp - qed - next - case u: (UINT b1 v1) - with p.IH have expr1: "expr e1 ep env cd st = expr (UINT b1 v1) ep env cd st" by simp - then show ?thesis - proof (cases "gas st > 0") - case True - then show ?thesis - proof (cases) - assume "b1 \ vbits" - with expr1 True - have "expr e1 ep env cd st=Normal ((KValue (createUInt b1 v1), Value (TUInt b1)),st)" by simp - moreover from u `b1 \ vbits` - have "v1 < 2^b1" and "v1 \ 0" using update_bounds_uint by auto - moreover from `b1 \ vbits` have "0 < b1" by auto - ultimately have r1: "expr e1 ep env cd st = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t v1), Value (TUInt b1)),st)" - by simp - thus ?thesis - proof (cases "eupdate e2") - case u2: (UINT b2 v2) - with p.IH have expr2: "expr e2 ep env cd st = expr (UINT b2 v2) ep env cd st" by simp - then show ?thesis - proof (cases) - let ?v="v1 + v2" - let ?x="?v mod 2 ^ max b1 b2" - assume "b2 \ vbits" - with expr2 True - have "expr e2 ep env cd st=Normal ((KValue (createUInt b2 v2), Value (TUInt b2)),st)" by simp - moreover from u2 `b2 \ vbits` - have "v2 < 2^b2" and "v2 \ 0" using update_bounds_uint by auto - moreover from `b2 \ vbits` have "0 < b2" by auto - ultimately have r2: "expr e2 ep env cd st = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t v2), Value (TUInt b2)),st)" - by simp - moreover have "add (TUInt b1) (TUInt b2) (ShowL\<^sub>i\<^sub>n\<^sub>t v1) (ShowL\<^sub>i\<^sub>n\<^sub>t v2) - = Some (createUInt (max b1 b2) ?v, TUInt (max b1 b2))" - using Read_ShowL_id add_def olift.simps(2)[of "(+)" b1 b2] by simp - ultimately have "expr (PLUS e1 e2) ep env cd st - = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TUInt (max b1 b2))),st)" using r1 True by simp - moreover have "expr (eupdate (PLUS e1 e2)) ep env cd st - = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TUInt (max b1 b2))),st)" - proof - - from `b1 \ vbits` `b2 \ vbits` - have "eupdate (PLUS e1 e2) = UINT (max b1 b2) ?x" using u u2 by simp - moreover have "expr (UINT (max b1 b2) ?x) ep env cd st - = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TUInt (max b1 b2))),st)" - proof - - from `b1 \ vbits` `b2 \ vbits` have "max b1 b2 \ vbits" using vbits_max by simp - with True have "expr (UINT (max b1 b2) ?x) ep env cd st - = Normal ((KValue (createUInt (max b1 b2) ?x), Value (TUInt (max b1 b2))),st)" by simp - moreover from `0 < b1` - have "?x < 2 ^ (max b1 b2)" by simp - moreover from `0 < b1` have "0 < max b1 b2" using max_def by simp - ultimately show ?thesis by simp - qed - ultimately show ?thesis by simp - qed - ultimately show ?thesis by simp - next - assume "\ b2 \ vbits" - with p u u2 show ?thesis by simp - qed - next - case i2: (INT b2 v2) - with p.IH have expr2: "expr e2 ep env cd st = expr (E.INT b2 v2) ep env cd st" by simp - then show ?thesis - proof (cases) - let ?v="v1 + v2" - assume "b2 \ vbits" - with expr2 True - have "expr e2 ep env cd st=Normal ((KValue (createSInt b2 v2), Value (TSInt b2)),st)" by simp - moreover from i2 `b2 \ vbits` - have "v2 < 2^(b2-1)" and "v2 \ -(2^(b2-1))" using update_bounds_int by auto - moreover from `b2 \ vbits` have "0 < b2" by auto - ultimately have r2: "expr e2 ep env cd st = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t v2), Value (TSInt b2)),st)" - using createSInt_id[of v2 b2] by simp - thus ?thesis - proof (cases) - assume "b10" - hence "createSInt b2 ?v = (ShowL\<^sub>i\<^sub>n\<^sub>t ?x)" using `b1i\<^sub>n\<^sub>t v1) (ShowL\<^sub>i\<^sub>n\<^sub>t v2) - = Some (createSInt b2 ?v, TSInt b2)" - using Read_ShowL_id add_def olift.simps(4)[of "(+)" b1 b2] `b1i\<^sub>n\<^sub>t ?x), Value (TSInt b2)),st)" using r1 r2 True by simp - moreover have "expr (eupdate (PLUS e1 e2)) ep env cd st - = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TSInt b2)),st)" - proof - - from `b1 \ vbits` `b2 \ vbits` `?v\0` `b1i\<^sub>n\<^sub>t ?x), Value (TSInt b2)),st)" - proof - - from `b2 \ vbits` True have "expr (E.INT b2 ?x) ep env cd st - = Normal ((KValue (createSInt b2 ?x), Value (TSInt b2)),st)" by simp - moreover from `0 < b2` have "?x < 2 ^ (b2 - 1)" using upper_bound2 by simp - ultimately show ?thesis using createSInt_id[of ?x "b2"] `0 < b2` by simp - qed - ultimately show ?thesis by simp - qed - ultimately show ?thesis by simp - next - let ?x="2^(b2 -1) - (-?v+2^(b2-1)-1) mod (2^b2) - 1" - assume "\ ?v\0" - hence "createSInt b2 ?v = (ShowL\<^sub>i\<^sub>n\<^sub>t ?x)" by simp - moreover have "add (TUInt b1) (TSInt b2) (ShowL\<^sub>i\<^sub>n\<^sub>t v1) (ShowL\<^sub>i\<^sub>n\<^sub>t v2) - = Some (createSInt b2 ?v, TSInt b2)" - using Read_ShowL_id add_def olift.simps(4)[of "(+)" b1 b2] `b1i\<^sub>n\<^sub>t ?x), Value (TSInt b2)),st)" using r1 r2 True by simp - moreover have "expr (eupdate (PLUS e1 e2)) ep env cd st - = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TSInt b2)),st)" - proof - - from `b1 \ vbits` `b2 \ vbits` `\?v\0` `b1i\<^sub>n\<^sub>t ?x), Value (TSInt b2)),st)" - proof - - from `b2 \ vbits` True have "expr (E.INT b2 ?x) ep env cd st - = Normal ((KValue (createSInt b2 ?x), Value (TSInt b2)),st)" by simp - moreover from `0 < b2` have "?x \ - (2 ^ (b2 - 1))" using upper_bound2 by simp - moreover have "2^(b2-1) - (-?v+2^(b2-1)-1) mod (2^b2) - 1 < 2 ^ (b2 - 1)" - by (simp add: algebra_simps flip: int_one_le_iff_zero_less) - ultimately show ?thesis using createSInt_id[of ?x b2] `0 < b2` by simp - qed - ultimately show ?thesis by simp - qed - ultimately show ?thesis by simp - qed - next - assume "\ b1 < b2" - with p u i2 show ?thesis by simp - qed - next - assume "\ b2 \ vbits" - with p u i2 show ?thesis by simp - qed - next - case (ADDRESS _) - with p u show ?thesis by simp - next - case (BALANCE _) - with p u show ?thesis by simp - next - case THIS - with p u show ?thesis by simp - next - case SENDER - with p u show ?thesis by simp - next - case VALUE - with p u show ?thesis by simp - next - case TRUE - with p u show ?thesis by simp - next - case FALSE - with p u show ?thesis by simp - next - case (LVAL _) - with p u show ?thesis by simp - next - case (PLUS _ _) - with p u show ?thesis by simp - next - case (MINUS _ _) - with p u show ?thesis by simp - next - case (EQUAL _ _) - with p u show ?thesis by simp - next - case (LESS _ _) - with p u show ?thesis by simp - next - case (AND _ _) - with p u show ?thesis by simp - next - case (OR _ _) - with p u show ?thesis by simp - next - case (NOT _) - with p u show ?thesis by simp - next - case (CALL x181 x182) - with p u show ?thesis by simp - next - case (ECALL x191 x192 x193 x194) - with p u show ?thesis by simp - qed - next - assume "\ b1 \ vbits" - with p u show ?thesis by simp - qed - next - case False - then show ?thesis using no_gas by simp - qed - next - case (ADDRESS x3) - with p show ?thesis by simp - next - case (BALANCE x4) - with p show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto - next - case THIS - with p show ?thesis by simp - next - case SENDER - with p show ?thesis by simp - next - case VALUE - with p show ?thesis by simp - next - case TRUE - with p show ?thesis by simp - next - case FALSE - with p show ?thesis by simp - next - case (LVAL x7) - with p show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto - next - case (PLUS x81 x82) - with p show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto - next - case (MINUS x91 x92) - with p show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto - next - case (EQUAL x101 x102) - with p show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto - next - case (LESS x111 x112) - with p show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto - next - case (AND x121 x122) - with p show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto - next - case (OR x131 x132) - with p show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto - next - case (NOT x131) - with p show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto - next - case (CALL x181 x182) - with p show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto - next - case (ECALL x191 x192 x193 x194) - with p show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto - qed -next - case m: (MINUS e1 e2) - show ?case - proof (cases "eupdate e1") - case i: (INT b1 v1) - with m.IH have expr1: "expr e1 ep env cd st = expr (E.INT b1 v1) ep env cd st" by simp - then show ?thesis - proof (cases "gas st > 0") - case True - show ?thesis - proof (cases) - assume "b1 \ vbits" - with expr1 True - have "expr e1 ep env cd st=Normal ((KValue (createSInt b1 v1), Value (TSInt b1)),st)" by simp - moreover from i `b1 \ vbits` - have "v1 < 2^(b1-1)" and "v1 \ -(2^(b1-1))" using update_bounds_int by auto - moreover from `b1 \ vbits` have "0 < b1" by auto - ultimately have r1: "expr e1 ep env cd st = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t v1), Value (TSInt b1)),st)" - using createSInt_id[of v1 b1] by simp - thus ?thesis - proof (cases "eupdate e2") - case i2: (INT b2 v2) - with m.IH have expr2: "expr e2 ep env cd st = expr (E.INT b2 v2) ep env cd st" by simp - then show ?thesis - proof (cases) - let ?v="v1 - v2" - assume "b2 \ vbits" - with expr2 True - have "expr e2 ep env cd st=Normal ((KValue (createSInt b2 v2), Value (TSInt b2)),st)" by simp - moreover from i2 `b2 \ vbits` - have "v2 < 2^(b2-1)" and "v2 \ -(2^(b2-1))" using update_bounds_int by auto - moreover from `b2 \ vbits` have "0 < b2" by auto - ultimately have r2: "expr e2 ep env cd st = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t v2), Value (TSInt b2)),st)" - using createSInt_id[of v2 b2] by simp - - from `b1 \ vbits` `b2 \ vbits` have - u_def: "eupdate (MINUS e1 e2) = - (let v = v1 - v2 - in if 0 \ v - then E.INT (max b1 b2) - (- (2 ^ (max b1 b2 - 1)) + (v + 2 ^ (max b1 b2 - 1)) mod 2 ^ max b1 b2) - else E.INT (max b1 b2) - (2 ^ (max b1 b2 - 1) - (- v + 2 ^ (max b1 b2 - 1) - 1) mod 2 ^ max b1 b2 - 1))" - using i i2 eupdate.simps(11)[of e1 e2] by simp - - show ?thesis - proof (cases) - let ?x="- (2 ^ (max b1 b2 - 1)) + (?v + 2 ^ (max b1 b2 - 1)) mod 2 ^ max b1 b2" - assume "?v\0" - hence "createSInt (max b1 b2) ?v = (ShowL\<^sub>i\<^sub>n\<^sub>t ?x)" by simp - moreover have "sub (TSInt b1) (TSInt b2) (ShowL\<^sub>i\<^sub>n\<^sub>t v1) (ShowL\<^sub>i\<^sub>n\<^sub>t v2) - = Some (createSInt (max b1 b2) ?v, TSInt (max b1 b2))" - using Read_ShowL_id sub_def olift.simps(1)[of "(-)" b1 b2] by simp - ultimately have "expr (MINUS e1 e2) ep env cd st - = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TSInt (max b1 b2))),st)" using r1 r2 True by simp - moreover have "expr (eupdate (MINUS e1 e2)) ep env cd st - = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TSInt (max b1 b2))),st)" - proof - - from u_def have "eupdate (MINUS e1 e2) = E.INT (max b1 b2) ?x" using `?v\0` by simp - moreover have "expr (E.INT (max b1 b2) ?x) ep env cd st - = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TSInt (max b1 b2))),st)" - proof - - from `b1 \ vbits` `b2 \ vbits` have "max b1 b2 \ vbits" using vbits_max by simp - with True have "expr (E.INT (max b1 b2) ?x) ep env cd st - = Normal ((KValue (createSInt (max b1 b2) ?x), Value (TSInt (max b1 b2))),st)" by simp - moreover from `0 < b1` - have "?x < 2 ^ (max b1 b2 - 1)" using upper_bound2 by simp - moreover from `0 < b1` have "0 < max b1 b2" using max_def by simp - ultimately show ?thesis using createSInt_id[of ?x "max b1 b2"] by simp - qed - ultimately show ?thesis by simp - qed - ultimately show ?thesis by simp - next - let ?x="2^(max b1 b2 -1) - (-?v+2^(max b1 b2-1)-1) mod (2^max b1 b2) - 1" - assume "\ ?v\0" - hence "createSInt (max b1 b2) ?v = (ShowL\<^sub>i\<^sub>n\<^sub>t ?x)" by simp - moreover have "sub (TSInt b1) (TSInt b2) (ShowL\<^sub>i\<^sub>n\<^sub>t v1) (ShowL\<^sub>i\<^sub>n\<^sub>t v2) - = Some (createSInt (max b1 b2) ?v, TSInt (max b1 b2))" - using Read_ShowL_id sub_def olift.simps(1)[of "(-)" b1 b2] by simp - ultimately have "expr (MINUS e1 e2) ep env cd st - = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TSInt (max b1 b2))),st)" using r1 r2 True by simp - moreover have "expr (eupdate (MINUS e1 e2)) ep env cd st - = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TSInt (max b1 b2))),st)" - proof - - from u_def have "eupdate (MINUS e1 e2) = E.INT (max b1 b2) ?x" using `\ ?v\0` by simp - moreover have "expr (E.INT (max b1 b2) ?x) ep env cd st - = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TSInt (max b1 b2))),st)" - proof - - from `b1 \ vbits` `b2 \ vbits` have "max b1 b2 \ vbits" using vbits_max by simp - with True have "expr (E.INT (max b1 b2) ?x) ep env cd st - = Normal ((KValue (createSInt (max b1 b2) ?x), Value (TSInt (max b1 b2))),st)" by simp - moreover from `0 < b1` - have "?x \ - (2 ^ (max b1 b2 - 1))" using lower_bound2[of "max b1 b2" ?v] by simp - moreover from `b1 > 0` have "2^(max b1 b2 -1) > (0::nat)" by simp - hence "2^(max b1 b2 -1) - (-?v+2^(max b1 b2-1)-1) mod (2^max b1 b2) - 1 < 2 ^ (max b1 b2 - 1)" - by (simp add: algebra_simps flip: int_one_le_iff_zero_less) - moreover from `0 < b1` have "0 < max b1 b2" using max_def by simp - ultimately show ?thesis using createSInt_id[of ?x "max b1 b2"] by simp - qed - ultimately show ?thesis by simp - qed - ultimately show ?thesis by simp - qed - next - assume "\ b2 \ vbits" - with m i i2 show ?thesis by simp - qed - next - case u: (UINT b2 v2) - with m.IH have expr2: "expr e2 ep env cd st = expr (UINT b2 v2) ep env cd st" by simp - then show ?thesis - proof (cases) - let ?v="v1 - v2" - assume "b2 \ vbits" - with expr2 True - have "expr e2 ep env cd st=Normal ((KValue (createUInt b2 v2), Value (TUInt b2)),st)" by simp - moreover from u `b2 \ vbits` - have "v2 < 2^b2" and "v2 \ 0" using update_bounds_uint by auto - moreover from `b2 \ vbits` have "0 < b2" by auto - ultimately have r2: "expr e2 ep env cd st = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t v2), Value (TUInt b2)),st)" - using createUInt_id[of v2 b2] by simp - thus ?thesis - proof (cases) - assume "b2 vbits` `b2 \ vbits` have - u_def: "eupdate (MINUS e1 e2) = - (let v = v1 - v2 - in if 0 \ v - then E.INT b1 (- (2 ^ (b1 - 1)) + (v + 2 ^ (b1 - 1)) mod 2 ^ b1) - else E.INT b1 (2 ^ (b1 - 1) - (- v + 2 ^ (b1 - 1) - 1) mod 2 ^ b1 - 1))" - using i u eupdate.simps(11)[of e1 e2] by simp - show ?thesis - proof (cases) - let ?x="- (2 ^ (b1 - 1)) + (?v + 2 ^ (b1 - 1)) mod 2 ^ b1" - assume "?v\0" - hence "createSInt b1 ?v = (ShowL\<^sub>i\<^sub>n\<^sub>t ?x)" using `b2i\<^sub>n\<^sub>t v1) (ShowL\<^sub>i\<^sub>n\<^sub>t v2) - = Some (createSInt b1 ?v, TSInt b1)" - using Read_ShowL_id sub_def olift.simps(3)[of "(-)" b1 b2] `b2i\<^sub>n\<^sub>t ?x), Value (TSInt b1)),st)" using r1 r2 True by simp - moreover have "expr (eupdate (MINUS e1 e2)) ep env cd st - = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TSInt b1)),st)" - proof - - from u_def have "eupdate (MINUS e1 e2) = E.INT b1 ?x" using `?v\0` by simp - moreover have "expr (E.INT b1 ?x) ep env cd st - = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TSInt b1)),st)" - proof - - from `b1 \ vbits` True have "expr (E.INT b1 ?x) ep env cd st - = Normal ((KValue (createSInt b1 ?x), Value (TSInt b1)),st)" by simp - moreover from `0 < b1` have "?x < 2 ^ (b1 - 1)" using upper_bound2 by simp - ultimately show ?thesis using createSInt_id[of ?x "b1"] `0 < b1` by simp - qed - ultimately show ?thesis by simp - qed - ultimately show ?thesis by simp - next - let ?x="2^(b1 -1) - (-?v+2^(b1-1)-1) mod (2^b1) - 1" - assume "\ ?v\0" - hence "createSInt b1 ?v = (ShowL\<^sub>i\<^sub>n\<^sub>t ?x)" by simp - moreover have "sub (TSInt b1) (TUInt b2) (ShowL\<^sub>i\<^sub>n\<^sub>t v1) (ShowL\<^sub>i\<^sub>n\<^sub>t v2) - = Some (createSInt b1 ?v, TSInt b1)" - using Read_ShowL_id sub_def olift.simps(3)[of "(-)" b1 b2] `b2i\<^sub>n\<^sub>t ?x), Value (TSInt b1)),st)" using r1 r2 True by simp - moreover have "expr (eupdate (MINUS e1 e2)) ep env cd st - = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TSInt b1)),st)" - proof - - from u_def have "eupdate (MINUS e1 e2) = E.INT b1 ?x" using `\ ?v\0` by simp - moreover have "expr (E.INT b1 ?x) ep env cd st - = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TSInt b1)),st)" - proof - - from `b1 \ vbits` True have "expr (E.INT b1 ?x) ep env cd st - = Normal ((KValue (createSInt b1 ?x), Value (TSInt b1)),st)" by simp - moreover from `0 < b1` have "?x \ - (2 ^ (b1 - 1))" using upper_bound2 by simp - moreover have "2^(b1-1) - (-?v+2^(b1-1)-1) mod (2^b1) - 1 < 2 ^ (b1 - 1)" - by (simp add: algebra_simps flip: int_one_le_iff_zero_less) - ultimately show ?thesis using createSInt_id[of ?x b1] `0 < b1` by simp - qed - ultimately show ?thesis by simp - qed - ultimately show ?thesis by simp - qed - next - assume "\ b2 < b1" - with m i u show ?thesis by simp - qed - next - assume "\ b2 \ vbits" - with m i u show ?thesis by simp - qed - next - case (ADDRESS _) - with m i show ?thesis by simp - next - case (BALANCE _) - with m i show ?thesis by simp - next - case THIS - with m i show ?thesis by simp - next - case SENDER - with m i show ?thesis by simp - next - case VALUE - with m i show ?thesis by simp - next - case TRUE - with m i show ?thesis by simp - next - case FALSE - with m i show ?thesis by simp - next - case (LVAL _) - with m i show ?thesis by simp - next - case (PLUS _ _) - with m i show ?thesis by simp - next - case (MINUS _ _) - with m i show ?thesis by simp - next - case (EQUAL _ _) - with m i show ?thesis by simp - next - case (LESS _ _) - with m i show ?thesis by simp - next - case (AND _ _) - with m i show ?thesis by simp - next - case (OR _ _) - with m i show ?thesis by simp - next - case (NOT _) - with m i show ?thesis by simp - next - case (CALL x181 x182) - with m i show ?thesis by simp - next - case (ECALL x191 x192 x193 x194) - with m i show ?thesis by simp - qed - next - assume "\ b1 \ vbits" - with m i show ?thesis by simp - qed - next - case False - then show ?thesis using no_gas by simp - qed - next - case u: (UINT b1 v1) - with m.IH have expr1: "expr e1 ep env cd st = expr (UINT b1 v1) ep env cd st" by simp - then show ?thesis - proof (cases "gas st > 0") - case True - show ?thesis - proof (cases) - assume "b1 \ vbits" - with expr1 True - have "expr e1 ep env cd st=Normal ((KValue (createUInt b1 v1), Value (TUInt b1)),st)" by simp - moreover from u `b1 \ vbits` - have "v1 < 2^b1" and "v1 \ 0" using update_bounds_uint by auto - moreover from `b1 \ vbits` have "0 < b1" by auto - ultimately have r1: "expr e1 ep env cd st = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t v1), Value (TUInt b1)),st)" - by simp - thus ?thesis - proof (cases "eupdate e2") - case u2: (UINT b2 v2) - with m.IH have expr2: "expr e2 ep env cd st = expr (UINT b2 v2) ep env cd st" by simp - then show ?thesis - proof (cases) - let ?v="v1 - v2" - let ?x="?v mod 2 ^ max b1 b2" - assume "b2 \ vbits" - with expr2 True - have "expr e2 ep env cd st=Normal ((KValue (createUInt b2 v2), Value (TUInt b2)),st)" by simp - moreover from u2 `b2 \ vbits` - have "v2 < 2^b2" and "v2 \ 0" using update_bounds_uint by auto - moreover from `b2 \ vbits` have "0 < b2" by auto - ultimately have r2: "expr e2 ep env cd st = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t v2), Value (TUInt b2)),st)" - by simp - moreover have "sub (TUInt b1) (TUInt b2) (ShowL\<^sub>i\<^sub>n\<^sub>t v1) (ShowL\<^sub>i\<^sub>n\<^sub>t v2) - = Some (createUInt (max b1 b2) ?v, TUInt (max b1 b2))" - using Read_ShowL_id sub_def olift.simps(2)[of "(-)" b1 b2] by simp - ultimately have "expr (MINUS e1 e2) ep env cd st - = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TUInt (max b1 b2))),st)" using r1 True by simp - moreover have "expr (eupdate (MINUS e1 e2)) ep env cd st - = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TUInt (max b1 b2))),st)" - proof - - from `b1 \ vbits` `b2 \ vbits` - have "eupdate (MINUS e1 e2) = UINT (max b1 b2) ?x" using u u2 by simp - moreover have "expr (UINT (max b1 b2) ?x) ep env cd st - = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TUInt (max b1 b2))),st)" - proof - - from `b1 \ vbits` `b2 \ vbits` have "max b1 b2 \ vbits" using vbits_max by simp - with True have "expr (UINT (max b1 b2) ?x) ep env cd st - = Normal ((KValue (createUInt (max b1 b2) ?x), Value (TUInt (max b1 b2))),st)" by simp - moreover from `0 < b1` - have "?x < 2 ^ (max b1 b2)" by simp - moreover from `0 < b1` have "0 < max b1 b2" using max_def by simp - ultimately show ?thesis by simp - qed - ultimately show ?thesis by simp - qed - ultimately show ?thesis by simp - next - assume "\ b2 \ vbits" - with m u u2 show ?thesis by simp - qed - next - case i: (INT b2 v2) - with m.IH have expr2: "expr e2 ep env cd st = expr (E.INT b2 v2) ep env cd st" by simp - then show ?thesis - proof (cases) - let ?v="v1 - v2" - assume "b2 \ vbits" - with expr2 True - have "expr e2 ep env cd st=Normal ((KValue (createSInt b2 v2), Value (TSInt b2)),st)" by simp - moreover from i `b2 \ vbits` - have "v2 < 2^(b2-1)" and "v2 \ -(2^(b2-1))" using update_bounds_int by auto - moreover from `b2 \ vbits` have "0 < b2" by auto - ultimately have r2: "expr e2 ep env cd st = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t v2), Value (TSInt b2)),st)" - using createSInt_id[of v2 b2] by simp - thus ?thesis - proof (cases) - assume "b1 vbits` `b2 \ vbits` have - u_def: "eupdate (MINUS e1 e2) = - (let v = v1 - v2 - in if 0 \ v - then E.INT b2 (- (2 ^ (b2 - 1)) + (v + 2 ^ (b2 - 1)) mod 2 ^ b2) - else E.INT b2 (2 ^ (b2 - 1) - (- v + 2 ^ (b2 - 1) - 1) mod 2 ^ b2 - 1))" - using u i by simp - show ?thesis - proof (cases) - let ?x="- (2 ^ (b2 - 1)) + (?v + 2 ^ (b2 - 1)) mod 2 ^ b2" - assume "?v\0" - hence "createSInt b2 ?v = (ShowL\<^sub>i\<^sub>n\<^sub>t ?x)" using `b1i\<^sub>n\<^sub>t v1) (ShowL\<^sub>i\<^sub>n\<^sub>t v2) - = Some (createSInt b2 ?v, TSInt b2)" - using Read_ShowL_id sub_def olift.simps(4)[of "(-)" b1 b2] `b1i\<^sub>n\<^sub>t ?x), Value (TSInt b2)),st)" using r1 r2 True by simp - moreover have "expr (eupdate (MINUS e1 e2)) ep env cd st - = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TSInt b2)),st)" - proof - - from u_def have "eupdate (MINUS e1 e2) = E.INT b2 ?x" using `?v\0` by simp - moreover have "expr (E.INT b2 ?x) ep env cd st - = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TSInt b2)),st)" - proof - - from `b2 \ vbits` True have "expr (E.INT b2 ?x) ep env cd st - = Normal ((KValue (createSInt b2 ?x), Value (TSInt b2)),st)" by simp - moreover from `0 < b2` have "?x < 2 ^ (b2 - 1)" using upper_bound2 by simp - ultimately show ?thesis using createSInt_id[of ?x "b2"] `0 < b2` by simp - qed - ultimately show ?thesis by simp - qed - ultimately show ?thesis by simp - next - let ?x="2^(b2 -1) - (-?v+2^(b2-1)-1) mod (2^b2) - 1" - assume "\ ?v\0" - hence "createSInt b2 ?v = (ShowL\<^sub>i\<^sub>n\<^sub>t ?x)" by simp - moreover have "sub (TUInt b1) (TSInt b2) (ShowL\<^sub>i\<^sub>n\<^sub>t v1) (ShowL\<^sub>i\<^sub>n\<^sub>t v2) - = Some (createSInt b2 ?v, TSInt b2)" - using Read_ShowL_id sub_def olift.simps(4)[of "(-)" b1 b2] `b1i\<^sub>n\<^sub>t ?x), Value (TSInt b2)),st)" using r1 r2 True by simp - moreover have "expr (eupdate (MINUS e1 e2)) ep env cd st - = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TSInt b2)),st)" - proof - - from u_def have "eupdate (MINUS e1 e2) = E.INT b2 ?x" using `\ ?v\0` by simp - moreover have "expr (E.INT b2 ?x) ep env cd st - = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TSInt b2)),st)" - proof - - from `b2 \ vbits` True have "expr (E.INT b2 ?x) ep env cd st - = Normal ((KValue (createSInt b2 ?x), Value (TSInt b2)),st)" by simp - moreover from `0 < b2` have "?x \ - (2 ^ (b2 - 1))" using upper_bound2 by simp - moreover have "2^(b2-1) - (-?v+2^(b2-1)-1) mod (2^b2) - 1 < 2 ^ (b2 - 1)" - by (simp add: algebra_simps flip: int_one_le_iff_zero_less) - ultimately show ?thesis using createSInt_id[of ?x b2] `0 < b2` by simp - qed - ultimately show ?thesis by simp - qed - ultimately show ?thesis by simp - qed - next - assume "\ b1 < b2" - with m u i show ?thesis by simp - qed - next - assume "\ b2 \ vbits" - with m u i show ?thesis by simp - qed - next - case (ADDRESS _) - with m u show ?thesis by simp - next - case (BALANCE _) - with m u show ?thesis by simp - next - case THIS - with m u show ?thesis by simp - next - case SENDER - with m u show ?thesis by simp - next - case VALUE - with m u show ?thesis by simp - next - case TRUE - with m u show ?thesis by simp - next - case FALSE - with m u show ?thesis by simp - next - case (LVAL _) - with m u show ?thesis by simp - next - case (PLUS _ _) - with m u show ?thesis by simp - next - case (MINUS _ _) - with m u show ?thesis by simp - next - case (EQUAL _ _) - with m u show ?thesis by simp - next - case (LESS _ _) - with m u show ?thesis by simp - next - case (AND _ _) - with m u show ?thesis by simp - next - case (OR _ _) - with m u show ?thesis by simp - next - case (NOT _) - with m u show ?thesis by simp - next - case (CALL x181 x182) - with m u show ?thesis by simp - next - case (ECALL x191 x192 x193 x194) - with m u show ?thesis by simp - qed - next - assume "\ b1 \ vbits" - with m u show ?thesis by simp - qed - next - case False - then show ?thesis using no_gas by simp - qed - next - case (ADDRESS x3) - with m show ?thesis by simp - next - case (BALANCE x4) - with m show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto - next - case THIS - with m show ?thesis by simp - next - case SENDER - with m show ?thesis by simp - next - case VALUE - with m show ?thesis by simp - next - case TRUE - with m show ?thesis by simp - next - case FALSE - with m show ?thesis by simp - next - case (LVAL x7) - with m show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto - next - case (PLUS x81 x82) - with m show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto - next - case (MINUS x91 x92) - with m show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto - next - case (EQUAL x101 x102) - with m show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto - next - case (LESS x111 x112) - with m show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto - next - case (AND x121 x122) - with m show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto - next - case (OR x131 x132) - with m show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto - next - case (NOT x131) - with m show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto - next - case (CALL x181 x182) - with m show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by simp - next - case (ECALL x191 x192 x193 x194) - with m show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by simp - qed -next - case e: (EQUAL e1 e2) - show ?case - proof (cases "eupdate e1") - case i: (INT b1 v1) - with e.IH have expr1: "expr e1 ep env cd st = expr (E.INT b1 v1) ep env cd st" by simp - then show ?thesis - proof (cases "gas st > 0") - case True - then show ?thesis - proof (cases) - assume "b1 \ vbits" - with expr1 True - have "expr e1 ep env cd st=Normal ((KValue (createSInt b1 v1), Value (TSInt b1)),st)" by simp - moreover from i `b1 \ vbits` - have "v1 < 2^(b1-1)" and "v1 \ -(2^(b1-1))" using update_bounds_int by auto - moreover from `b1 \ vbits` have "0 < b1" by auto - ultimately have r1: "expr e1 ep env cd st = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t v1), Value (TSInt b1)),st)" - using createSInt_id[of v1 b1] by simp - thus ?thesis - proof (cases "eupdate e2") - case i2: (INT b2 v2) - with e.IH have expr2: "expr e2 ep env cd st = expr (E.INT b2 v2) ep env cd st" by simp - then show ?thesis - proof (cases) - assume "b2 \ vbits" - with expr2 True - have "expr e2 ep env cd st=Normal ((KValue (createSInt b2 v2), Value (TSInt b2)),st)" by simp - moreover from i2 `b2 \ vbits` - have "v2 < 2^(b2-1)" and "v2 \ -(2^(b2-1))" using update_bounds_int by auto - moreover from `b2 \ vbits` have "0 < b2" by auto - ultimately have r2: "expr e2 ep env cd st = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t v2), Value (TSInt b2)),st)" - using createSInt_id[of v2 b2] by simp - with r1 True have "expr (EQUAL e1 e2) ep env cd st= - Normal ((KValue (createBool ((ReadL\<^sub>i\<^sub>n\<^sub>t (ShowL\<^sub>i\<^sub>n\<^sub>t v1))=((ReadL\<^sub>i\<^sub>n\<^sub>t (ShowL\<^sub>i\<^sub>n\<^sub>t v2))))), Value TBool),st)" - using equal_def plift.simps(1)[of "(=)"] by simp - hence "expr (EQUAL e1 e2) ep env cd st=Normal ((KValue (createBool (v1=v2)), Value TBool),st)" - using Read_ShowL_id by simp - with `b1 \ vbits` `b2 \ vbits` True show ?thesis using i i2 by simp - next - assume "\ b2 \ vbits" - with e i i2 show ?thesis by simp - qed - next - case u: (UINT b2 v2) - with e.IH have expr2: "expr e2 ep env cd st = expr (UINT b2 v2) ep env cd st" by simp - then show ?thesis - proof (cases) - assume "b2 \ vbits" - with expr2 True - have "expr e2 ep env cd st=Normal ((KValue (createUInt b2 v2), Value (TUInt b2)),st)" by simp - moreover from u `b2 \ vbits` - have "v2 < 2^b2" and "v2 \ 0" using update_bounds_uint by auto - moreover from `b2 \ vbits` have "0 < b2" by auto - ultimately have r2: "expr e2 ep env cd st = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t v2), Value (TUInt b2)),st)" - using createUInt_id[of v2 b2] by simp - thus ?thesis - proof (cases) - assume "b2i\<^sub>n\<^sub>t (ShowL\<^sub>i\<^sub>n\<^sub>t v1))=((ReadL\<^sub>i\<^sub>n\<^sub>t (ShowL\<^sub>i\<^sub>n\<^sub>t v2))))), Value TBool),st)" - using equal_def plift.simps(3)[of "(=)"] by simp - hence "expr (EQUAL e1 e2) ep env cd st=Normal ((KValue (createBool (v1=v2)), Value TBool),st)" - using Read_ShowL_id by simp - with `b1 \ vbits` `b2 \ vbits` `b2 b2 < b1" - with e i u show ?thesis by simp - qed - next - assume "\ b2 \ vbits" - with e i u show ?thesis by simp - qed - next - case (ADDRESS _) - with e i show ?thesis by simp - next - case (BALANCE _) - with e i show ?thesis by simp - next - case THIS - with e i show ?thesis by simp - next - case SENDER - with e i show ?thesis by simp - next - case VALUE - with e i show ?thesis by simp - next - case TRUE - with e i show ?thesis by simp - next - case FALSE - with e i show ?thesis by simp - next - case (LVAL _) - with e i show ?thesis by simp - next - case (PLUS _ _) - with e i show ?thesis by simp - next - case (MINUS _ _) - with e i show ?thesis by simp - next - case (EQUAL _ _) - with e i show ?thesis by simp - next - case (LESS _ _) - with e i show ?thesis by simp - next - case (AND _ _) - with e i show ?thesis by simp - next - case (OR _ _) - with e i show ?thesis by simp - next - case (NOT _) - with e i show ?thesis by simp - next - case (CALL x181 x182) - with e i show ?thesis by simp - next - case (ECALL x191 x192 x193 x194) - with e i show ?thesis by simp - qed - next - assume "\ b1 \ vbits" - with e i show ?thesis by simp - qed - next - case False - then show ?thesis using no_gas by simp - qed - next - case u: (UINT b1 v1) - with e.IH have expr1: "expr e1 ep env cd st = expr (UINT b1 v1) ep env cd st" by simp - then show ?thesis - proof (cases "gas st > 0") - case True - then show ?thesis - proof (cases) - assume "b1 \ vbits" - with expr1 True - have "expr e1 ep env cd st=Normal ((KValue (createUInt b1 v1), Value (TUInt b1)),st)" by simp - moreover from u `b1 \ vbits` - have "v1 < 2^b1" and "v1 \ 0" using update_bounds_uint by auto - moreover from `b1 \ vbits` have "0 < b1" by auto - ultimately have r1: "expr e1 ep env cd st = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t v1), Value (TUInt b1)),st)" - by simp - thus ?thesis - proof (cases "eupdate e2") - case u2: (UINT b2 v2) - with e.IH have expr2: "expr e2 ep env cd st = expr (UINT b2 v2) ep env cd st" by simp - then show ?thesis - proof (cases) - assume "b2 \ vbits" - with expr2 True - have "expr e2 ep env cd st=Normal ((KValue (createUInt b2 v2), Value (TUInt b2)),st)" by simp - moreover from u2 `b2 \ vbits` - have "v2 < 2^b2" and "v2 \ 0" using update_bounds_uint by auto - moreover from `b2 \ vbits` have "0 < b2" by auto - ultimately have r2: "expr e2 ep env cd st = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t v2), Value (TUInt b2)),st)" - by simp - with r1 True have "expr (EQUAL e1 e2) ep env cd st= - Normal ((KValue (createBool ((ReadL\<^sub>i\<^sub>n\<^sub>t (ShowL\<^sub>i\<^sub>n\<^sub>t v1))=((ReadL\<^sub>i\<^sub>n\<^sub>t (ShowL\<^sub>i\<^sub>n\<^sub>t v2))))), Value TBool),st)" - using equal_def plift.simps(2)[of "(=)"] by simp - hence "expr (EQUAL e1 e2) ep env cd st=Normal ((KValue (createBool (v1=v2)), Value TBool),st)" - using Read_ShowL_id by simp - with `b1 \ vbits` `b2 \ vbits` show ?thesis using u u2 True by simp - next - assume "\ b2 \ vbits" - with e u u2 show ?thesis by simp - qed - next - case i: (INT b2 v2) - with e.IH have expr2: "expr e2 ep env cd st = expr (E.INT b2 v2) ep env cd st" by simp - then show ?thesis - proof (cases) - let ?v="v1 + v2" - assume "b2 \ vbits" - with expr2 True - have "expr e2 ep env cd st=Normal ((KValue (createSInt b2 v2), Value (TSInt b2)),st)" by simp - moreover from i `b2 \ vbits` - have "v2 < 2^(b2-1)" and "v2 \ -(2^(b2-1))" using update_bounds_int by auto - moreover from `b2 \ vbits` have "0 < b2" by auto - ultimately have r2: "expr e2 ep env cd st = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t v2), Value (TSInt b2)),st)" - using createSInt_id[of v2 b2] by simp - thus ?thesis - proof (cases) - assume "b1i\<^sub>n\<^sub>t (ShowL\<^sub>i\<^sub>n\<^sub>t v1))=((ReadL\<^sub>i\<^sub>n\<^sub>t (ShowL\<^sub>i\<^sub>n\<^sub>t v2))))), Value TBool),st)" - using equal_def plift.simps(4)[of "(=)"] by simp - hence "expr (EQUAL e1 e2) ep env cd st=Normal ((KValue (createBool (v1=v2)), Value TBool),st)" - using Read_ShowL_id by simp - with `b1 \ vbits` `b2 \ vbits` `b1 b1 < b2" - with e u i show ?thesis by simp - qed - next - assume "\ b2 \ vbits" - with e u i show ?thesis by simp - qed - next - case (ADDRESS _) - with e u show ?thesis by simp - next - case (BALANCE _) - with e u show ?thesis by simp - next - case THIS - with e u show ?thesis by simp - next - case SENDER - with e u show ?thesis by simp - next - case VALUE - with e u show ?thesis by simp - next - case TRUE - with e u show ?thesis by simp - next - case FALSE - with e u show ?thesis by simp - next - case (LVAL _) - with e u show ?thesis by simp - next - case (PLUS _ _) - with e u show ?thesis by simp - next - case (MINUS _ _) - with e u show ?thesis by simp - next - case (EQUAL _ _) - with e u show ?thesis by simp - next - case (LESS _ _) - with e u show ?thesis by simp - next - case (AND _ _) - with e u show ?thesis by simp - next - case (OR _ _) - with e u show ?thesis by simp - next - case (NOT _) - with e u show ?thesis by simp - next - case (CALL x181 x182) - with e u show ?thesis by simp - next - case (ECALL x191 x192 x193 x194) - with e u show ?thesis by simp - qed - next - assume "\ b1 \ vbits" - with e u show ?thesis by simp - qed - next - case False - then show ?thesis using no_gas by simp - qed - next - case (ADDRESS x3) - with e show ?thesis by simp - next - case (BALANCE x4) - with e show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto - next - case THIS - with e show ?thesis by simp - next - case SENDER - with e show ?thesis by simp - next - case VALUE - with e show ?thesis by simp - next - case TRUE - with e show ?thesis by simp - next - case FALSE - with e show ?thesis by simp - next - case (LVAL x7) - with e show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto - next - case (PLUS x81 x82) - with e show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto - next - case (MINUS x91 x92) - with e show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto - next - case (EQUAL x101 x102) - with e show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto - next - case (LESS x111 x112) - with e show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto - next - case (AND x121 x122) - with e show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto - next - case (OR x131 x132) - with e show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto - next - case (NOT x131) - with e show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto - next - case (CALL x181 x182) - with e show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by simp - next - case (ECALL x191 x192 x193 x194) - with e show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by simp - qed -next - case l: (LESS e1 e2) - show ?case - proof (cases "eupdate e1") - case i: (INT b1 v1) - with l.IH have expr1: "expr e1 ep env cd st = expr (E.INT b1 v1) ep env cd st" by simp - then show ?thesis - proof (cases "gas st > 0") - case True - then show ?thesis - proof (cases) - assume "b1 \ vbits" - with expr1 True - have "expr e1 ep env cd st=Normal ((KValue (createSInt b1 v1), Value (TSInt b1)),st)" by simp - moreover from i `b1 \ vbits` - have "v1 < 2^(b1-1)" and "v1 \ -(2^(b1-1))" using update_bounds_int by auto - moreover from `b1 \ vbits` have "0 < b1" by auto - ultimately have r1: "expr e1 ep env cd st = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t v1), Value (TSInt b1)),st)" - using createSInt_id[of v1 b1] by simp - thus ?thesis - proof (cases "eupdate e2") - case i2: (INT b2 v2) - with l.IH have expr2: "expr e2 ep env cd st = expr (E.INT b2 v2) ep env cd st" by simp - then show ?thesis - proof (cases) - assume "b2 \ vbits" - with expr2 True - have "expr e2 ep env cd st=Normal ((KValue (createSInt b2 v2), Value (TSInt b2)),st)" by simp - moreover from i2 `b2 \ vbits` - have "v2 < 2^(b2-1)" and "v2 \ -(2^(b2-1))" using update_bounds_int by auto - moreover from `b2 \ vbits` have "0 < b2" by auto - ultimately have r2: "expr e2 ep env cd st = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t v2), Value (TSInt b2)),st)" - using createSInt_id[of v2 b2] by simp - with r1 True have "expr (LESS e1 e2) ep env cd st= - Normal ((KValue (createBool ((ReadL\<^sub>i\<^sub>n\<^sub>t (ShowL\<^sub>i\<^sub>n\<^sub>t v1))<((ReadL\<^sub>i\<^sub>n\<^sub>t (ShowL\<^sub>i\<^sub>n\<^sub>t v2))))), Value TBool),st)" - using less_def plift.simps(1)[of "(<)"] by simp - hence "expr (LESS e1 e2) ep env cd st=Normal ((KValue (createBool (v1 vbits` `b2 \ vbits` show ?thesis using i i2 True by simp - next - assume "\ b2 \ vbits" - with l i i2 show ?thesis by simp - qed - next - case u: (UINT b2 v2) - with l.IH have expr2: "expr e2 ep env cd st = expr (UINT b2 v2) ep env cd st" by simp - then show ?thesis - proof (cases) - assume "b2 \ vbits" - with expr2 True - have "expr e2 ep env cd st=Normal ((KValue (createUInt b2 v2), Value (TUInt b2)),st)" by simp - moreover from u `b2 \ vbits` - have "v2 < 2^b2" and "v2 \ 0" using update_bounds_uint by auto - moreover from `b2 \ vbits` have "0 < b2" by auto - ultimately have r2: "expr e2 ep env cd st = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t v2), Value (TUInt b2)),st)" - using createUInt_id[of v2 b2] by simp - thus ?thesis - proof (cases) - assume "b2i\<^sub>n\<^sub>t (ShowL\<^sub>i\<^sub>n\<^sub>t v1))<((ReadL\<^sub>i\<^sub>n\<^sub>t (ShowL\<^sub>i\<^sub>n\<^sub>t v2))))), Value TBool),st)" - using less_def plift.simps(3)[of "(<)"] by simp - hence "expr (LESS e1 e2) ep env cd st=Normal ((KValue (createBool (v1 vbits` `b2 \ vbits` `b2 b2 < b1" - with l i u show ?thesis by simp - qed - next - assume "\ b2 \ vbits" - with l i u show ?thesis by simp - qed - next - case (ADDRESS _) - with l i show ?thesis by simp - next - case (BALANCE _) - with l i show ?thesis by simp - next - case THIS - with l i show ?thesis by simp - next - case SENDER - with l i show ?thesis by simp - next - case VALUE - with l i show ?thesis by simp - next - case TRUE - with l i show ?thesis by simp - next - case FALSE - with l i show ?thesis by simp - next - case (LVAL _) - with l i show ?thesis by simp - next - case (PLUS _ _) - with l i show ?thesis by simp - next - case (MINUS _ _) - with l i show ?thesis by simp - next - case (EQUAL _ _) - with l i show ?thesis by simp - next - case (LESS _ _) - with l i show ?thesis by simp - next - case (AND _ _) - with l i show ?thesis by simp - next - case (OR _ _) - with l i show ?thesis by simp - next - case (NOT _) - with l i show ?thesis by simp - next - case (CALL x181 x182) - with l i show ?thesis by simp - next - case (ECALL x191 x192 x193 x194) - with l i show ?thesis by simp - qed - next - assume "\ b1 \ vbits" - with l i show ?thesis by simp - qed - next - case False - then show ?thesis using no_gas by simp - qed - next - case u: (UINT b1 v1) - with l.IH have expr1: "expr e1 ep env cd st = expr (UINT b1 v1) ep env cd st" by simp - then show ?thesis - proof (cases "gas st > 0") - case True - then show ?thesis - proof (cases) - assume "b1 \ vbits" - with expr1 True - have "expr e1 ep env cd st=Normal ((KValue (createUInt b1 v1), Value (TUInt b1)),st)" by simp - moreover from u `b1 \ vbits` - have "v1 < 2^b1" and "v1 \ 0" using update_bounds_uint by auto - moreover from `b1 \ vbits` have "0 < b1" by auto - ultimately have r1: "expr e1 ep env cd st = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t v1), Value (TUInt b1)),st)" - by simp - thus ?thesis - proof (cases "eupdate e2") - case u2: (UINT b2 v2) - with l.IH have expr2: "expr e2 ep env cd st = expr (UINT b2 v2) ep env cd st" by simp - then show ?thesis - proof (cases) - assume "b2 \ vbits" - with expr2 True - have "expr e2 ep env cd st=Normal ((KValue (createUInt b2 v2), Value (TUInt b2)),st)" by simp - moreover from u2 `b2 \ vbits` - have "v2 < 2^b2" and "v2 \ 0" using update_bounds_uint by auto - moreover from `b2 \ vbits` have "0 < b2" by auto - ultimately have r2: "expr e2 ep env cd st = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t v2), Value (TUInt b2)),st)" - by simp - with r1 True have "expr (LESS e1 e2) ep env cd st=Normal ((KValue (createBool ((ReadL\<^sub>i\<^sub>n\<^sub>t (ShowL\<^sub>i\<^sub>n\<^sub>t v1))<((ReadL\<^sub>i\<^sub>n\<^sub>t (ShowL\<^sub>i\<^sub>n\<^sub>t v2))))), Value TBool),st)" using less_def plift.simps(2)[of "(<)"] by simp - hence "expr (LESS e1 e2) ep env cd st=Normal ((KValue (createBool (v1 vbits` `b2 \ vbits` show ?thesis using u u2 True by simp - next - assume "\ b2 \ vbits" - with l u u2 show ?thesis by simp - qed - next - case i: (INT b2 v2) - with l.IH have expr2: "expr e2 ep env cd st = expr (E.INT b2 v2) ep env cd st" by simp - then show ?thesis - proof (cases) - let ?v="v1 + v2" - assume "b2 \ vbits" - with expr2 True - have "expr e2 ep env cd st=Normal ((KValue (createSInt b2 v2), Value (TSInt b2)),st)" by simp - moreover from i `b2 \ vbits` - have "v2 < 2^(b2-1)" and "v2 \ -(2^(b2-1))" using update_bounds_int by auto - moreover from `b2 \ vbits` have "0 < b2" by auto - ultimately have r2: "expr e2 ep env cd st = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t v2), Value (TSInt b2)),st)" - using createSInt_id[of v2 b2] by simp - thus ?thesis - proof (cases) - assume "b1i\<^sub>n\<^sub>t (ShowL\<^sub>i\<^sub>n\<^sub>t v1))<((ReadL\<^sub>i\<^sub>n\<^sub>t (ShowL\<^sub>i\<^sub>n\<^sub>t v2))))), Value TBool),st)" - using less_def plift.simps(4)[of "(<)"] by simp - hence "expr (LESS e1 e2) ep env cd st=Normal ((KValue (createBool (v1 vbits` `b2 \ vbits` `b1 b1 < b2" - with l u i show ?thesis by simp - qed - next - assume "\ b2 \ vbits" - with l u i show ?thesis by simp - qed - next - case (ADDRESS _) - with l u show ?thesis by simp - next - case (BALANCE _) - with l u show ?thesis by simp - next - case THIS - with l u show ?thesis by simp - next - case SENDER - with l u show ?thesis by simp - next - case VALUE - with l u show ?thesis by simp - next - case TRUE - with l u show ?thesis by simp - next - case FALSE - with l u show ?thesis by simp - next - case (LVAL _) - with l u show ?thesis by simp - next - case (PLUS _ _) - with l u show ?thesis by simp - next - case (MINUS _ _) - with l u show ?thesis by simp - next - case (EQUAL _ _) - with l u show ?thesis by simp - next - case (LESS _ _) - with l u show ?thesis by simp - next - case (AND _ _) - with l u show ?thesis by simp - next - case (OR _ _) - with l u show ?thesis by simp - next - case (NOT _) - with l u show ?thesis by simp - next - case (CALL x181 x182) - with l u show ?thesis by simp - next - case (ECALL x191 x192 x193 x194) - with l u show ?thesis by simp - qed - next - assume "\ b1 \ vbits" - with l u show ?thesis by simp - qed - next - case False - then show ?thesis using no_gas by simp - qed - next - case (ADDRESS x3) - with l show ?thesis by simp - next - case (BALANCE x4) - with l show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto - next - case THIS - with l show ?thesis by simp - next - case SENDER - with l show ?thesis by simp - next - case VALUE - with l show ?thesis by simp - next - case TRUE - with l show ?thesis by simp - next - case FALSE - with l show ?thesis by simp - next - case (LVAL x7) - with l show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto - next - case (PLUS x81 x82) - with l show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto - next - case (MINUS x91 x92) - with l show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto - next - case (EQUAL x101 x102) - with l show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto - next - case (LESS x111 x112) - with l show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto - next - case (AND x121 x122) - with l show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto - next - case (OR x131 x132) - with l show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto - next - case (NOT x131) - with l show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto - next - case (CALL x181 x182) - with l show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by simp - next - case (ECALL x191 x192 x193 x194) - with l show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by simp - qed -next - case a: (AND e1 e2) - show ?case - proof (cases "eupdate e1") - case (INT x11 x12) - with a show ?thesis by simp - next - case (UINT x21 x22) - with a show ?thesis by simp - next - case (ADDRESS x3) - with a show ?thesis by simp - next - case (BALANCE x4) - with a show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto - next - case THIS - with a show ?thesis by simp - next - case SENDER - with a show ?thesis by simp - next - case VALUE - with a show ?thesis by simp - next - case t: TRUE - show ?thesis - proof (cases "eupdate e2") - case (INT x11 x12) - with a t show ?thesis by simp - next - case (UINT x21 x22) - with a t show ?thesis by simp - next - case (ADDRESS x3) - with a t show ?thesis by simp - next - case (BALANCE x4) - with a t show ?thesis by simp - next - case THIS - with a t show ?thesis by simp - next - case SENDER - with a t show ?thesis by simp - next - case VALUE - with a t show ?thesis by simp - next - case TRUE - with a t show ?thesis by simp - next - case FALSE - with a t show ?thesis by simp - next - case (LVAL x7) - with a t show ?thesis by simp - next - case (PLUS x81 x82) - with a t show ?thesis by simp - next - case (MINUS x91 x92) - with a t show ?thesis by simp - next - case (EQUAL x101 x102) - with a t show ?thesis by simp - next - case (LESS x111 x112) - with a t show ?thesis by simp - next - case (AND x121 x122) - with a t show ?thesis by simp - next - case (OR x131 x132) - with a t show ?thesis by simp - next - case (NOT x131) - with a t show ?thesis by simp - next - case (CALL x181 x182) - with a t show ?thesis by simp - next - case (ECALL x191 x192 x193 x194) - with a t show ?thesis by simp - qed - next - case f: FALSE - show ?thesis - proof (cases "eupdate e2") - case (INT b v) - with a f show ?thesis by simp - next - case (UINT b v) - with a f show ?thesis by simp - next - case (ADDRESS x3) - with a f show ?thesis by simp - next - case (BALANCE x4) - with a f show ?thesis by simp - next - case THIS - with a f show ?thesis by simp - next - case SENDER - with a f show ?thesis by simp - next - case VALUE - with a f show ?thesis by simp - next - case TRUE - with a f show ?thesis by simp - next - case FALSE - with a f show ?thesis by simp - next - case (LVAL x7) - with a f show ?thesis by simp - next - case (PLUS x81 x82) - with a f show ?thesis by simp - next - case (MINUS x91 x92) - with a f show ?thesis by simp - next - case (EQUAL x101 x102) - with a f show ?thesis by simp - next - case (LESS x111 x112) - with a f show ?thesis by simp - next - case (AND x121 x122) - with a f show ?thesis by simp - next - case (OR x131 x132) - with a f show ?thesis by simp - next - case (NOT x131) - with a f show ?thesis by simp - next - case (CALL x181 x182) - with a f show ?thesis by simp - next - case (ECALL x191 x192 x193 x194) - with a f show ?thesis by simp - qed - next - case (LVAL x7) - with a show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto - next - case p: (PLUS x81 x82) - with a show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto - next - case (MINUS x91 x92) - with a show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto - next - case (EQUAL x101 x102) - with a show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto - next - case (LESS x111 x112) - with a show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto - next - case (AND x121 x122) - with a show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto - next - case (OR x131 x132) - with a show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto - next - case (NOT x131) - with a show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto - next - case (CALL x181 x182) - with a show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by simp - next - case (ECALL x191 x192 x193 x194) - with a show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by simp - qed -next - case o: (OR e1 e2) - show ?case - proof (cases "eupdate e1") - case (INT x11 x12) - with o show ?thesis by simp - next - case (UINT x21 x22) - with o show ?thesis by simp - next - case (ADDRESS x3) - with o show ?thesis by simp - next - case (BALANCE x4) - with o show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto - next - case THIS - with o show ?thesis by simp - next - case SENDER - with o show ?thesis by simp - next - case VALUE - with o show ?thesis by simp - next - case t: TRUE - show ?thesis - proof (cases "eupdate e2") - case (INT x11 x12) - with o t show ?thesis by simp - next - case (UINT x21 x22) - with o t show ?thesis by simp - next - case (ADDRESS x3) - with o t show ?thesis by simp - next - case (BALANCE x4) - with o t show ?thesis by simp - next - case THIS - with o t show ?thesis by simp - next - case SENDER - with o t show ?thesis by simp - next - case VALUE - with o t show ?thesis by simp - next - case TRUE - with o t show ?thesis by simp - next - case FALSE - with o t show ?thesis by simp - next - case (LVAL x7) - with o t show ?thesis by simp - next - case (PLUS x81 x82) - with o t show ?thesis by simp - next - case (MINUS x91 x92) - with o t show ?thesis by simp - next - case (EQUAL x101 x102) - with o t show ?thesis by simp - next - case (LESS x111 x112) - with o t show ?thesis by simp - next - case (AND x121 x122) - with o t show ?thesis by simp - next - case (OR x131 x132) - with o t show ?thesis by simp - next - case (NOT x131) - with o t show ?thesis by simp - next - case (CALL x181 x182) - with o t show ?thesis by simp - next - case (ECALL x191 x192 x193 x194) - with o t show ?thesis by simp - qed - next - case f: FALSE - show ?thesis - proof (cases "eupdate e2") - case (INT b v) - with o f show ?thesis by simp - next - case (UINT b v) - with o f show ?thesis by simp - next - case (ADDRESS x3) - with o f show ?thesis by simp - next - case (BALANCE x4) - with o f show ?thesis by simp - next - case THIS - with o f show ?thesis by simp - next - case SENDER - with o f show ?thesis by simp - next - case VALUE - with o f show ?thesis by simp - next - case TRUE - with o f show ?thesis by simp - next - case FALSE - with o f show ?thesis by simp - next - case (LVAL x7) - with o f show ?thesis by simp - next - case (PLUS x81 x82) - with o f show ?thesis by simp - next - case (MINUS x91 x92) - with o f show ?thesis by simp - next - case (EQUAL x101 x102) - with o f show ?thesis by simp - next - case (LESS x111 x112) - with o f show ?thesis by simp - next - case (AND x121 x122) - with o f show ?thesis by simp - next - case (OR x131 x132) - with o f show ?thesis by simp - next - case (NOT x131) - with o f show ?thesis by simp - next - case (CALL x181 x182) - with o f show ?thesis by simp - next - case (ECALL x191 x192 x193 x194) - with o f show ?thesis by simp - qed - next - case (LVAL x7) - with o show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto - next - case p: (PLUS x81 x82) - with o show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto - next - case (MINUS x91 x92) - with o show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto - next - case (EQUAL x101 x102) - with o show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto - next - case (LESS x111 x112) - with o show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto - next - case (AND x121 x122) - with o show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto - next - case (OR x131 x132) - with o show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto - next - case (NOT x131) - with o show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto - next - case (CALL x181 x182) - with o show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by simp - next - case (ECALL x191 x192 x193 x194) - with o show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by simp - qed -next - case o: (NOT e) - show ?case - proof (cases "eupdate e") - case (INT x11 x12) - with o show ?thesis by simp - next - case (UINT x21 x22) - with o show ?thesis by simp - next - case (ADDRESS x3) - with o show ?thesis by simp - next - case (BALANCE x4) - with o show ?thesis by simp - next - case THIS - with o show ?thesis by simp - next - case SENDER - with o show ?thesis by simp - next - case VALUE - with o show ?thesis by simp - next - case t: TRUE - with o show ?thesis by simp - next - case f: FALSE - with o show ?thesis by simp - next - case (LVAL x7) - with o show ?thesis by simp - next - case p: (PLUS x81 x82) - with o show ?thesis by simp - next - case (MINUS x91 x92) - with o show ?thesis by simp - next - case (EQUAL x101 x102) - with o show ?thesis by simp - next - case (LESS x111 x112) - with o show ?thesis by simp - next - case (AND x121 x122) - with o show ?thesis by simp - next - case (OR x131 x132) - with o show ?thesis by simp - next - case (NOT x131) - with o show ?thesis by simp - next - case (CALL x181 x182) - with o show ?thesis by simp - next - case (ECALL x191 x192 x193 x194) - with o show ?thesis by simp - qed -next - case (CALL x181 x182) - show ?case by simp -next - case (ECALL x191 x192 x193 x194) - show ?case by simp -qed - -end \ No newline at end of file +section\Constant Folding\ +theory Constant_Folding +imports + Solidity_Main +begin +text\ + The following function optimizes expressions w.r.t. gas consumption. +\ +primrec eupdate :: "E \ E" +and lupdate :: "L \ L" +where + "lupdate (Id i) = Id i" +| "lupdate (Ref i xs) = Ref i (map eupdate xs)" +| "eupdate (E.INT b v) = + (if (b\vbits) + then if v \ 0 + then E.INT b (-(2^(b-1)) + (v+2^(b-1)) mod (2^b)) + else E.INT b (2^(b-1) - (-v+2^(b-1)-1) mod (2^b) - 1) + else E.INT b v)" +| "eupdate (UINT b v) = (if (b\vbits) then UINT b (v mod (2^b)) else UINT b v)" +| "eupdate (ADDRESS a) = ADDRESS a" +| "eupdate (BALANCE a) = BALANCE a" +| "eupdate THIS = THIS" +| "eupdate SENDER = SENDER" +| "eupdate VALUE = VALUE" +| "eupdate TRUE = TRUE" +| "eupdate FALSE = FALSE" +| "eupdate (LVAL l) = LVAL (lupdate l)" +| "eupdate (PLUS ex1 ex2) = + (case (eupdate ex1) of + E.INT b1 v1 \ + if b1 \ vbits + then (case (eupdate ex2) of + E.INT b2 v2 \ + if b2\vbits + then let v=v1+v2 in + if v \ 0 + then E.INT (max b1 b2) (-(2^((max b1 b2)-1)) + (v+2^((max b1 b2)-1)) mod (2^(max b1 b2))) + else E.INT (max b1 b2) (2^((max b1 b2)-1) - (-v+2^((max b1 b2)-1)-1) mod (2^(max b1 b2)) - 1) + else (PLUS (E.INT b1 v1) (E.INT b2 v2)) + | UINT b2 v2 \ + if b2\vbits \ b2 < b1 + then let v=v1+v2 in + if v \ 0 + then E.INT b1 (-(2^(b1-1)) + (v+2^(b1-1)) mod (2^b1)) + else E.INT b1 (2^(b1-1) - (-v+2^(b1-1)-1) mod (2^b1) - 1) + else PLUS (E.INT b1 v1) (UINT b2 v2) + | _ \ PLUS (E.INT b1 v1) (eupdate ex2)) + else PLUS (E.INT b1 v1) (eupdate ex2) + | UINT b1 v1 \ + if b1 \ vbits + then (case (eupdate ex2) of + UINT b2 v2 \ + if b2 \ vbits + then UINT (max b1 b2) ((v1 + v2) mod (2^(max b1 b2))) + else (PLUS (UINT b1 v1) (UINT b2 v2)) + | E.INT b2 v2 \ + if b2\vbits \ b1 < b2 + then let v=v1+v2 in + if v \ 0 + then E.INT b2 (-(2^(b2-1)) + (v+2^(b2-1)) mod (2^b2)) + else E.INT b2 (2^(b2-1) - (-v+2^(b2-1)-1) mod (2^b2) - 1) + else PLUS (UINT b1 v1) (E.INT b2 v2) + | _ \ PLUS (UINT b1 v1) (eupdate ex2)) + else PLUS (UINT b1 v1) (eupdate ex2) + | _ \ PLUS (eupdate ex1) (eupdate ex2))" +| "eupdate (MINUS ex1 ex2) = + (case (eupdate ex1) of + E.INT b1 v1 \ + if b1 \ vbits + then (case (eupdate ex2) of + E.INT b2 v2 \ + if b2\vbits + then let v=v1-v2 in + if v \ 0 + then E.INT (max b1 b2) (-(2^((max b1 b2)-1)) + (v+2^((max b1 b2)-1)) mod (2^(max b1 b2))) + else E.INT (max b1 b2) (2^((max b1 b2)-1) - (-v+2^((max b1 b2)-1)-1) mod (2^(max b1 b2)) - 1) + else (MINUS (E.INT b1 v1) (E.INT b2 v2)) + | UINT b2 v2 \ + if b2\vbits \ b2 < b1 + then let v=v1-v2 in + if v \ 0 + then E.INT b1 (-(2^(b1-1)) + (v+2^(b1-1)) mod (2^b1)) + else E.INT b1 (2^(b1-1) - (-v+2^(b1-1)-1) mod (2^b1) - 1) + else MINUS (E.INT b1 v1) (UINT b2 v2) + | _ \ MINUS (E.INT b1 v1) (eupdate ex2)) + else MINUS (E.INT b1 v1) (eupdate ex2) + | UINT b1 v1 \ + if b1 \ vbits + then (case (eupdate ex2) of + UINT b2 v2 \ + if b2 \ vbits + then UINT (max b1 b2) ((v1 - v2) mod (2^(max b1 b2))) + else (MINUS (UINT b1 v1) (UINT b2 v2)) + | E.INT b2 v2 \ + if b2\vbits \ b1 < b2 + then let v=v1-v2 in + if v \ 0 + then E.INT b2 (-(2^(b2-1)) + (v+2^(b2-1)) mod (2^b2)) + else E.INT b2 (2^(b2-1) - (-v+2^(b2-1)-1) mod (2^b2) - 1) + else MINUS (UINT b1 v1) (E.INT b2 v2) + | _ \ MINUS (UINT b1 v1) (eupdate ex2)) + else MINUS (UINT b1 v1) (eupdate ex2) + | _ \ MINUS (eupdate ex1) (eupdate ex2))" +| "eupdate (EQUAL ex1 ex2) = + (case (eupdate ex1) of + E.INT b1 v1 \ + if b1 \ vbits + then (case (eupdate ex2) of + E.INT b2 v2 \ + if b2\vbits + then if v1 = v2 + then TRUE + else FALSE + else EQUAL (E.INT b1 v1) (E.INT b2 v2) + | UINT b2 v2 \ + if b2\vbits \ b2 < b1 + then if v1 = v2 + then TRUE + else FALSE + else EQUAL (E.INT b1 v1) (UINT b2 v2) + | _ \ EQUAL (E.INT b1 v1) (eupdate ex2)) + else EQUAL (E.INT b1 v1) (eupdate ex2) + | UINT b1 v1 \ + if b1 \ vbits + then (case (eupdate ex2) of + UINT b2 v2 \ + if b2 \ vbits + then if v1 = v2 + then TRUE + else FALSE + else EQUAL (E.INT b1 v1) (UINT b2 v2) + | E.INT b2 v2 \ + if b2\vbits \ b1 < b2 + then if v1 = v2 + then TRUE + else FALSE + else EQUAL (UINT b1 v1) (E.INT b2 v2) + | _ \ EQUAL (UINT b1 v1) (eupdate ex2)) + else EQUAL (UINT b1 v1) (eupdate ex2) + | _ \ EQUAL (eupdate ex1) (eupdate ex2))" +| "eupdate (LESS ex1 ex2) = + (case (eupdate ex1) of + E.INT b1 v1 \ + if b1 \ vbits + then (case (eupdate ex2) of + E.INT b2 v2 \ + if b2\vbits + then if v1 < v2 + then TRUE + else FALSE + else LESS (E.INT b1 v1) (E.INT b2 v2) + | UINT b2 v2 \ + if b2\vbits \ b2 < b1 + then if v1 < v2 + then TRUE + else FALSE + else LESS (E.INT b1 v1) (UINT b2 v2) + | _ \ LESS (E.INT b1 v1) (eupdate ex2)) + else LESS (E.INT b1 v1) (eupdate ex2) + | UINT b1 v1 \ + if b1 \ vbits + then (case (eupdate ex2) of + UINT b2 v2 \ + if b2 \ vbits + then if v1 < v2 + then TRUE + else FALSE + else LESS (E.INT b1 v1) (UINT b2 v2) + | E.INT b2 v2 \ + if b2\vbits \ b1 < b2 + then if v1 < v2 + then TRUE + else FALSE + else LESS (UINT b1 v1) (E.INT b2 v2) + | _ \ LESS (UINT b1 v1) (eupdate ex2)) + else LESS (UINT b1 v1) (eupdate ex2) + | _ \ LESS (eupdate ex1) (eupdate ex2))" +| "eupdate (AND ex1 ex2) = + (case (eupdate ex1) of + TRUE \ (case (eupdate ex2) of + TRUE \ TRUE + | FALSE \ FALSE + | _ \ AND TRUE (eupdate ex2)) + | FALSE \ (case (eupdate ex2) of + TRUE \ FALSE + | FALSE \ FALSE + | _ \ AND FALSE (eupdate ex2)) + | _ \ AND (eupdate ex1) (eupdate ex2))" +| "eupdate (OR ex1 ex2) = + (case (eupdate ex1) of + TRUE \ (case (eupdate ex2) of + TRUE \ TRUE + | FALSE \ TRUE + | _ \ OR TRUE (eupdate ex2)) + | FALSE \ (case (eupdate ex2) of + TRUE \ TRUE + | FALSE \ FALSE + | _ \ OR FALSE (eupdate ex2)) + | _ \ OR (eupdate ex1) (eupdate ex2))" +| "eupdate (NOT ex1) = + (case (eupdate ex1) of + TRUE \ FALSE + | FALSE \ TRUE + | _ \ NOT (eupdate ex1))" +| "eupdate (CALL i xs) = CALL i xs" +| "eupdate (ECALL e i xs) = ECALL e i xs" +| "eupdate CONTRACTS = CONTRACTS" + +lemma "eupdate (UINT 8 250) + =UINT 8 250" + by(simp) +lemma "eupdate (UINT 8 500) + = UINT 8 244" + by(simp) +lemma "eupdate (E.INT 8 (-100)) + = E.INT 8 (- 100)" + by(simp) +lemma "eupdate (E.INT 8 (-150)) + = E.INT 8 106" + by(simp) +lemma "eupdate (PLUS (UINT 8 100) (UINT 8 100)) + = UINT 8 200" + by(simp) +lemma "eupdate (PLUS (UINT 8 257) (UINT 16 100)) + = UINT 16 101" + by(simp) +lemma "eupdate (PLUS (E.INT 8 100) (UINT 8 250)) + = PLUS (E.INT 8 100) (UINT 8 250)" + by(simp) +lemma "eupdate (PLUS (E.INT 8 250) (UINT 8 500)) + = PLUS (E.INT 8 (- 6)) (UINT 8 244)" + by(simp) +lemma "eupdate (PLUS (E.INT 16 250) (UINT 8 500)) + = E.INT 16 494" + by(simp) +lemma "eupdate (EQUAL (UINT 16 250) (UINT 8 250)) + = TRUE" + by(simp) +lemma "eupdate (EQUAL (E.INT 16 100) (UINT 8 100)) + = TRUE" + by(simp) +lemma "eupdate (EQUAL (E.INT 8 100) (UINT 8 100)) + = EQUAL (E.INT 8 100) (UINT 8 100)" + by(simp) + +lemma update_bounds_int: + assumes "eupdate ex = (E.INT b v)" and "b\vbits" + shows "(v < 2^(b-1)) \ v \ -(2^(b-1))" +proof (cases ex) + case (INT b' v') + then show ?thesis + proof cases + assume "b'\vbits" + show ?thesis + proof cases + let ?x="-(2^(b'-1)) + (v'+2^(b'-1)) mod 2^b'" + assume "v'\0" + with `b'\vbits` have "eupdate (E.INT b' v') = E.INT b' ?x" by simp + with assms have "b=b'" and "v=?x" using INT by (simp,simp) + moreover from `b'\vbits` have "b'>0" by auto + hence "?x < 2 ^(b'-1)" using upper_bound2[of b' "(v' + 2 ^ (b' - 1)) mod 2^b'"] by simp + moreover have "?x \ -(2^(b'-1))" by simp + ultimately show ?thesis by simp + next + let ?x="2^(b'-1) - (-v'+2^(b'-1)-1) mod (2^b') - 1" + assume "\v'\0" + with `b'\vbits` have "eupdate (E.INT b' v') = E.INT b' ?x" by simp + with assms have "b=b'" and "v=?x" using INT by (simp,simp) + moreover have "(-v'+2^(b'-1)-1) mod (2^b')\0" by simp + hence "?x < 2 ^(b'-1)" by arith + moreover from `b'\vbits` have "b'>0" by auto + hence "?x \ -(2^(b'-1))" using lower_bound2[of b' v'] by simp + ultimately show ?thesis by simp + qed + next + assume "\ b'\vbits" + with assms show ?thesis using INT by simp + qed +next + case (UINT b' v') + with assms show ?thesis + proof cases + assume "b'\vbits" + with assms show ?thesis using UINT by simp + next + assume "\ b'\vbits" + with assms show ?thesis using UINT by simp + qed +next + case (ADDRESS x3) + with assms show ?thesis by simp +next + case (BALANCE x4) + with assms show ?thesis by simp +next + case THIS + with assms show ?thesis by simp +next + case SENDER + with assms show ?thesis by simp +next + case VALUE + with assms show ?thesis by simp +next + case TRUE + with assms show ?thesis by simp +next + case FALSE + with assms show ?thesis by simp +next + case (LVAL x7) + with assms show ?thesis by simp +next + case p: (PLUS e1 e2) + show ?thesis + proof (cases "eupdate e1") + case i: (INT b1 v1) + show ?thesis + proof cases + assume "b1\vbits" + show ?thesis + proof (cases "eupdate e2") + case i2: (INT b2 v2) + then show ?thesis + proof cases + let ?v="v1+v2" + assume "b2\vbits" + show ?thesis + proof cases + let ?x="-(2^((max b1 b2)-1)) + (?v+2^((max b1 b2)-1)) mod 2^(max b1 b2)" + assume "?v\0" + with `b1\vbits` `b2\vbits` i i2 have "eupdate (PLUS e1 e2) = E.INT (max b1 b2) ?x" by simp + with assms have "b=max b1 b2" and "v=?x" using p by (simp,simp) + moreover from `b1\vbits` have "max b1 b2>0" by auto + hence "?x < 2 ^(max b1 b2 - 1)" + using upper_bound2[of "max b1 b2" "(?v + 2 ^ (max b1 b2 - 1)) mod 2^max b1 b2"] by simp + moreover have "?x \ -(2^(max b1 b2-1))" by simp + ultimately show ?thesis by simp + next + let ?x="2^((max b1 b2)-1) - (-?v+2^((max b1 b2)-1)-1) mod (2^(max b1 b2)) - 1" + assume "\?v\0" + with `b1\vbits` `b2\vbits` i i2 have "eupdate (PLUS e1 e2) = E.INT (max b1 b2) ?x" by simp + with assms have "b=max b1 b2" and "v=?x" using p by (simp,simp) + moreover have "(-?v+2^(max b1 b2-1)-1) mod (2^max b1 b2)\0" by simp + hence "?x < 2 ^(max b1 b2-1)" by arith + moreover from `b1\vbits` have "max b1 b2>0" by auto + hence "?x \ -(2^(max b1 b2-1))" using lower_bound2[of "max b1 b2" ?v] by simp + ultimately show ?thesis by simp + qed + next + assume "b2\vbits" + with p i i2 `b1\vbits` show ?thesis using assms by simp + qed + next + case u: (UINT b2 v2) + then show ?thesis + proof cases + let ?v="v1+v2" + assume "b2\vbits" + show ?thesis + proof cases + assume "b20" + with `b1\vbits` `b2\vbits` `b2vbits` have "b1>0" by auto + hence "?x < 2 ^(b1 - 1)" using upper_bound2[of b1] by simp + moreover have "?x \ -(2^(b1-1))" by simp + ultimately show ?thesis by simp + next + let ?x="2^(b1-1) - (-?v+2^(b1-1)-1) mod (2^b1) - 1" + assume "\?v\0" + with `b1\vbits` `b2\vbits` `b20" by simp + hence "?x < 2 ^(b1-1)" by arith + moreover from `b1\vbits` have "b1>0" by auto + hence "?x \ -(2^(b1-1))" using lower_bound2[of b1 ?v] by simp + ultimately show ?thesis by simp + qed + next + assume "\ b2vbits` show ?thesis using assms by simp + qed + next + assume "b2\vbits" + with p i u `b1\vbits` show ?thesis using assms by simp + qed + next + case (ADDRESS x3) + with p i `b1\vbits` show ?thesis using assms by simp + next + case (BALANCE x4) + with p i `b1\vbits` show ?thesis using assms by simp + next + case THIS + with p i `b1\vbits` show ?thesis using assms by simp + next + case SENDER + with p i `b1\vbits` show ?thesis using assms by simp + next + case VALUE + with p i `b1\vbits` show ?thesis using assms by simp + next + case TRUE + with p i `b1\vbits` show ?thesis using assms by simp + next + case FALSE + with p i `b1\vbits` show ?thesis using assms by simp + next + case (LVAL x7) + with p i `b1\vbits` show ?thesis using assms by simp + next + case (PLUS x81 x82) + with p i `b1\vbits` show ?thesis using assms by simp + next + case (MINUS x91 x92) + with p i `b1\vbits` show ?thesis using assms by simp + next + case (EQUAL x101 x102) + with p i `b1\vbits` show ?thesis using assms by simp + next + case (LESS x111 x112) + with p i `b1\vbits` show ?thesis using assms by simp + next + case (AND x121 x122) + with p i `b1\vbits` show ?thesis using assms by simp + next + case (OR x131 x132) + with p i `b1\vbits` show ?thesis using assms by simp + next + case (NOT x131) + with p i `b1\vbits` show ?thesis using assms by simp + next + case (CALL x181 x182) + with p i `b1\vbits` show ?thesis using assms by simp + next + case (ECALL x191 x192 x193) + with p i `b1\vbits` show ?thesis using assms by simp + next + case CONTRACTS + with p i `b1\vbits` show ?thesis using assms by simp + qed + next + assume "\ b1\vbits" + with p i show ?thesis using assms by simp + qed + next + case u: (UINT b1 v1) + show ?thesis + proof cases + assume "b1\vbits" + show ?thesis + proof (cases "eupdate e2") + case i: (INT b2 v2) + then show ?thesis + proof cases + let ?v="v1+v2" + assume "b2\vbits" + show ?thesis + proof cases + assume "b10" + with `b1\vbits` `b2\vbits` `b1vbits` have "b2>0" by auto + hence "?x < 2 ^(b2 - 1)" using upper_bound2[of b2] by simp + moreover have "?x \ -(2^(b2-1))" by simp + ultimately show ?thesis by simp + next + let ?x="2^(b2-1) - (-?v+2^(b2-1)-1) mod (2^b2) - 1" + assume "\?v\0" + with `b1\vbits` `b2\vbits` `b10" by simp + hence "?x < 2 ^(b2-1)" by arith + moreover from `b2\vbits` have "b2>0" by auto + hence "?x \ -(2^(b2-1))" using lower_bound2[of b2 ?v] by simp + ultimately show ?thesis by simp + qed + next + assume "\ b1vbits` show ?thesis using assms by simp + qed + next + assume "b2\vbits" + with p i u `b1\vbits` show ?thesis using assms by simp + qed + next + case u2: (UINT b2 v2) + then show ?thesis + proof cases + assume "b2\vbits" + with `b1\vbits` u u2 p show ?thesis using assms by simp + next + assume "\b2\vbits" + with p u u2 `b1\vbits` show ?thesis using assms by simp + qed + next + case (ADDRESS x3) + with p u `b1\vbits` show ?thesis using assms by simp + next + case (BALANCE x4) + with p u `b1\vbits` show ?thesis using assms by simp + next + case THIS + with p u `b1\vbits` show ?thesis using assms by simp + next + case SENDER + with p u `b1\vbits` show ?thesis using assms by simp + next + case VALUE + with p u `b1\vbits` show ?thesis using assms by simp + next + case TRUE + with p u `b1\vbits` show ?thesis using assms by simp + next + case FALSE + with p u `b1\vbits` show ?thesis using assms by simp + next + case (LVAL x7) + with p u `b1\vbits` show ?thesis using assms by simp + next + case (PLUS x81 x82) + with p u `b1\vbits` show ?thesis using assms by simp + next + case (MINUS x91 x92) + with p u `b1\vbits` show ?thesis using assms by simp + next + case (EQUAL x101 x102) + with p u `b1\vbits` show ?thesis using assms by simp + next + case (LESS x111 x112) + with p u `b1\vbits` show ?thesis using assms by simp + next + case (AND x121 x122) + with p u `b1\vbits` show ?thesis using assms by simp + next + case (OR x131 x132) + with p u `b1\vbits` show ?thesis using assms by simp + next + case (NOT x131) + with p u `b1\vbits` show ?thesis using assms by simp + next + case (CALL x181 x182) + with p u `b1\vbits` show ?thesis using assms by simp + next + case (ECALL x191 x192 x193) + with p u `b1\vbits` show ?thesis using assms by simp + next + case CONTRACTS + with p u `b1\vbits` show ?thesis using assms by simp + qed + next + assume "\ b1\vbits" + with p u show ?thesis using assms by simp + qed + next + case (ADDRESS x3) + with p show ?thesis using assms by simp + next + case (BALANCE x4) + with p show ?thesis using assms by simp + next + case THIS + with p show ?thesis using assms by simp + next + case SENDER + with p show ?thesis using assms by simp + next + case VALUE + with p show ?thesis using assms by simp + next + case TRUE + with p show ?thesis using assms by simp + next + case FALSE + with p show ?thesis using assms by simp + next + case (LVAL x7) + with p show ?thesis using assms by simp + next + case (PLUS x81 x82) + with p show ?thesis using assms by simp + next + case (MINUS x91 x92) + with p show ?thesis using assms by simp + next + case (EQUAL x101 x102) + with p show ?thesis using assms by simp + next + case (LESS x111 x112) + with p show ?thesis using assms by simp + next + case (AND x121 x122) + with p show ?thesis using assms by simp + next + case (OR x131 x132) + with p show ?thesis using assms by simp + next + case (NOT x131) + with p show ?thesis using assms by simp + next + case (CALL x181 x182) + with p show ?thesis using assms by simp + next + case (ECALL x191 x192 x193) + with p show ?thesis using assms by simp + next + case CONTRACTS + with p show ?thesis using assms by simp + qed +next + case m: (MINUS e1 e2) + show ?thesis + proof (cases "eupdate e1") + case i: (INT b1 v1) + with m show ?thesis + proof cases + assume "b1\vbits" + show ?thesis + proof (cases "eupdate e2") + case i2: (INT b2 v2) + then show ?thesis + proof cases + let ?v="v1-v2" + assume "b2\vbits" + with `b1 \ vbits` have + u_def: "eupdate (MINUS e1 e2) = + (let v = v1 - v2 + in if 0 \ v + then E.INT (max b1 b2) + (- (2 ^ (max b1 b2 - 1)) + (v + 2 ^ (max b1 b2 - 1)) mod 2 ^ max b1 b2) + else E.INT (max b1 b2) + (2 ^ (max b1 b2 - 1) - (- v + 2 ^ (max b1 b2 - 1) - 1) mod 2 ^ max b1 b2 - 1))" + using i i2 eupdate.simps(11)[of e1 e2] by simp + show ?thesis + proof cases + let ?x="-(2^((max b1 b2)-1)) + (?v+2^((max b1 b2)-1)) mod 2^(max b1 b2)" + assume "?v\0" + with u_def have "eupdate (MINUS e1 e2) = E.INT (max b1 b2) ?x" by simp + with assms have "b=max b1 b2" and "v=?x" using m by (simp,simp) + moreover from `b1\vbits` have "max b1 b2>0" by auto + hence "?x < 2 ^(max b1 b2 - 1)" + using upper_bound2[of "max b1 b2" "(?v + 2 ^ (max b1 b2 - 1)) mod 2^max b1 b2"] by simp + moreover have "?x \ -(2^(max b1 b2-1))" by simp + ultimately show ?thesis by simp + next + let ?x="2^((max b1 b2)-1) - (-?v+2^((max b1 b2)-1)-1) mod (2^(max b1 b2)) - 1" + assume "\?v\0" + with u_def have "eupdate (MINUS e1 e2) = E.INT (max b1 b2) ?x" using u_def by simp + with assms have "b=max b1 b2" and "v=?x" using m by (simp,simp) + moreover have "(-?v+2^(max b1 b2-1)-1) mod (2^max b1 b2)\0" by simp + hence "?x < 2 ^(max b1 b2-1)" by arith + moreover from `b1\vbits` have "max b1 b2>0" by auto + hence "?x \ -(2^(max b1 b2-1))" using lower_bound2[of "max b1 b2" ?v] by simp + ultimately show ?thesis by simp + qed + next + assume "b2\vbits" + with m i i2 `b1\vbits` show ?thesis using assms by simp + qed + next + case u: (UINT b2 v2) + then show ?thesis + proof cases + let ?v="v1-v2" + assume "b2\vbits" + show ?thesis + proof cases + assume "b2 vbits` `b2 \ vbits` have + u_def: "eupdate (MINUS e1 e2) = + (let v = v1 - v2 + in if 0 \ v + then E.INT b1 (- (2 ^ (b1 - 1)) + (v + 2 ^ (b1 - 1)) mod 2 ^ b1) + else E.INT b1 (2 ^ (b1 - 1) - (- v + 2 ^ (b1 - 1) - 1) mod 2 ^ b1 - 1))" + using i u eupdate.simps(11)[of e1 e2] by simp + then show ?thesis + proof cases + let ?x="(-(2^(b1-1)) + (?v+2^(b1-1)) mod (2^b1))" + assume "?v\0" + with u_def have "eupdate (MINUS e1 e2) = E.INT b1 ?x" by simp + with assms have "b=b1" and "v=?x" using m by (simp,simp) + moreover from `b1\vbits` have "b1>0" by auto + hence "?x < 2 ^(b1 - 1)" using upper_bound2[of b1] by simp + moreover have "?x \ -(2^(b1-1))" by simp + ultimately show ?thesis by simp + next + let ?x="2^(b1-1) - (-?v+2^(b1-1)-1) mod (2^b1) - 1" + assume "\?v\0" + with u_def have "eupdate (MINUS e1 e2) = E.INT b1 ?x" by simp + with assms have "b=b1" and "v=?x" using m i u by (simp,simp) + moreover have "(-?v+2^(b1-1)-1) mod 2^b1\0" by simp + hence "?x < 2 ^(b1-1)" by arith + moreover from `b1\vbits` have "b1>0" by auto + hence "?x \ -(2^(b1-1))" using lower_bound2[of b1 ?v] by simp + ultimately show ?thesis by simp + qed + next + assume "\ b2vbits` show ?thesis using assms by simp + qed + next + assume "b2\vbits" + with m i u `b1\vbits` show ?thesis using assms by simp + qed + next + case (ADDRESS x3) + with m i `b1\vbits` show ?thesis using assms by simp + next + case (BALANCE x4) + with m i `b1\vbits` show ?thesis using assms by simp + next + case THIS + with m i `b1\vbits` show ?thesis using assms by simp + next + case SENDER + with m i `b1\vbits` show ?thesis using assms by simp + next + case VALUE + with m i `b1\vbits` show ?thesis using assms by simp + next + case TRUE + with m i `b1\vbits` show ?thesis using assms by simp + next + case FALSE + with m i `b1\vbits` show ?thesis using assms by simp + next + case (LVAL x7) + with m i `b1\vbits` show ?thesis using assms by simp + next + case (PLUS x81 x82) + with m i `b1\vbits` show ?thesis using assms by simp + next + case (MINUS x91 x92) + with m i `b1\vbits` show ?thesis using assms by simp + next + case (EQUAL x101 x102) + with m i `b1\vbits` show ?thesis using assms by simp + next + case (LESS x111 x112) + with m i `b1\vbits` show ?thesis using assms by simp + next + case (AND x121 x122) + with m i `b1\vbits` show ?thesis using assms by simp + next + case (OR x131 x132) + with m i `b1\vbits` show ?thesis using assms by simp + next + case (NOT x131) + with m i `b1\vbits` show ?thesis using assms by simp + next + case (CALL x181 x182) + with m i `b1\vbits` show ?thesis using assms by simp + next + case (ECALL x191 x192 x193) + with m i `b1\vbits` show ?thesis using assms by simp + next + case CONTRACTS + with m i `b1\vbits` show ?thesis using assms by simp + qed + next + assume "\ b1\vbits" + with m i show ?thesis using assms by simp + qed + next + case u: (UINT b1 v1) + show ?thesis + proof cases + assume "b1\vbits" + show ?thesis + proof (cases "eupdate e2") + case i: (INT b2 v2) + then show ?thesis + proof cases + let ?v="v1-v2" + assume "b2\vbits" + show ?thesis + proof cases + assume "b1 vbits` `b2 \ vbits` have + u_def: "eupdate (MINUS e1 e2) = + (let v = v1 - v2 + in if 0 \ v + then E.INT b2 (- (2 ^ (b2 - 1)) + (v + 2 ^ (b2 - 1)) mod 2 ^ b2) + else E.INT b2 (2 ^ (b2 - 1) - (- v + 2 ^ (b2 - 1) - 1) mod 2 ^ b2 - 1))" + using i u eupdate.simps(11)[of e1 e2] by simp + then show ?thesis + proof cases + let ?x="(-(2^(b2-1)) + (?v+2^(b2-1)) mod (2^b2))" + assume "?v\0" + with u_def have "eupdate (MINUS e1 e2) = E.INT b2 ?x" by simp + with assms have "b=b2" and "v=?x" using m by (simp,simp) + moreover from `b2\vbits` have "b2>0" by auto + hence "?x < 2 ^(b2 - 1)" using upper_bound2[of b2] by simp + moreover have "?x \ -(2^(b2-1))" by simp + ultimately show ?thesis by simp + next + let ?x="2^(b2-1) - (-?v+2^(b2-1)-1) mod (2^b2) - 1" + assume "\?v\0" + with u_def have "eupdate (MINUS e1 e2) = E.INT b2 ?x" by simp + with assms have "b=b2" and "v=?x" using m i u by (simp,simp) + moreover have "(-?v+2^(b2-1)-1) mod 2^b2\0" by simp + hence "?x < 2 ^(b2-1)" by arith + moreover from `b2\vbits` have "b2>0" by auto + hence "?x \ -(2^(b2-1))" using lower_bound2[of b2 ?v] by simp + ultimately show ?thesis by simp + qed + next + assume "\ b1vbits` show ?thesis using assms by simp + qed + next + assume "b2\vbits" + with m i u `b1\vbits` show ?thesis using assms by simp + qed + next + case u2: (UINT b2 v2) + then show ?thesis + proof cases + assume "b2\vbits" + with `b1\vbits` u u2 m show ?thesis using assms by simp + next + assume "\b2\vbits" + with m u u2 `b1\vbits` show ?thesis using assms by simp + qed + next + case (ADDRESS x3) + with m u `b1\vbits` show ?thesis using assms by simp + next + case (BALANCE x4) + with m u `b1\vbits` show ?thesis using assms by simp + next + case THIS + with m u `b1\vbits` show ?thesis using assms by simp + next + case SENDER + with m u `b1\vbits` show ?thesis using assms by simp + next + case VALUE + with m u `b1\vbits` show ?thesis using assms by simp + next + case TRUE + with m u `b1\vbits` show ?thesis using assms by simp + next + case FALSE + with m u `b1\vbits` show ?thesis using assms by simp + next + case (LVAL x7) + with m u `b1\vbits` show ?thesis using assms by simp + next + case (PLUS x81 x82) + with m u `b1\vbits` show ?thesis using assms by simp + next + case (MINUS x91 x92) + with m u `b1\vbits` show ?thesis using assms by simp + next + case (EQUAL x101 x102) + with m u `b1\vbits` show ?thesis using assms by simp + next + case (LESS x111 x112) + with m u `b1\vbits` show ?thesis using assms by simp + next + case (AND x121 x122) + with m u `b1\vbits` show ?thesis using assms by simp + next + case (OR x131 x132) + with m u `b1\vbits` show ?thesis using assms by simp + next + case (NOT x131) + with m u `b1\vbits` show ?thesis using assms by simp + next + case (CALL x181 x182) + with m u `b1\vbits` show ?thesis using assms by simp + next + case (ECALL x191 x192 x193) + with m u `b1\vbits` show ?thesis using assms by simp + next + case CONTRACTS + with m u `b1\vbits` show ?thesis using assms by simp + qed + next + assume "\ b1\vbits" + with m u show ?thesis using assms by simp + qed + next + case (ADDRESS x3) + with m show ?thesis using assms by simp + next + case (BALANCE x4) + with m show ?thesis using assms by simp + next + case THIS + with m show ?thesis using assms by simp + next + case SENDER + with m show ?thesis using assms by simp + next + case VALUE + with m show ?thesis using assms by simp + next + case TRUE + with m show ?thesis using assms by simp + next + case FALSE + with m show ?thesis using assms by simp + next + case (LVAL x7) + with m show ?thesis using assms by simp + next + case (PLUS x81 x82) + with m show ?thesis using assms by simp + next + case (MINUS x91 x92) + with m show ?thesis using assms by simp + next + case (EQUAL x101 x102) + with m show ?thesis using assms by simp + next + case (LESS x111 x112) + with m show ?thesis using assms by simp + next + case (AND x121 x122) + with m show ?thesis using assms by simp + next + case (OR x131 x132) + with m show ?thesis using assms by simp + next + case (NOT x131) + with m show ?thesis using assms by simp + next + case (CALL x181 x182) + with m show ?thesis using assms by simp + next + case (ECALL x191 x192 x193) + with m show ?thesis using assms by simp + next + case CONTRACTS + with m show ?thesis using assms by simp + qed +next + case e: (EQUAL e1 e2) + show ?thesis + proof (cases "eupdate e1") + case i: (INT b1 v1) + show ?thesis + proof cases + assume "b1\vbits" + show ?thesis + proof (cases "eupdate e2") + case i2: (INT b2 v2) + then show ?thesis + proof cases + assume "b2\vbits" + show ?thesis + proof cases + assume "v1=v2" + with assms show ?thesis using e i i2 `b1\vbits` `b2\vbits` by simp + next + assume "\ v1=v2" + with assms show ?thesis using e i i2 `b1\vbits` `b2\vbits` by simp + qed + next + assume "b2\vbits" + with e i i2 `b1\vbits` show ?thesis using assms by simp + qed + next + case u: (UINT b2 v2) + then show ?thesis + proof cases + assume "b2\vbits" + show ?thesis + proof cases + assume "b2vbits` `b2\vbits` `b2 v1=v2" + with assms show ?thesis using e i u `b1\vbits` `b2\vbits` `b2 b2vbits` show ?thesis using assms by simp + qed + next + assume "b2\vbits" + with e i u `b1\vbits` show ?thesis using assms by simp + qed + next + case (ADDRESS x3) + with e i `b1\vbits` show ?thesis using assms by simp + next + case (BALANCE x4) + with e i `b1\vbits` show ?thesis using assms by simp + next + case THIS + with e i `b1\vbits` show ?thesis using assms by simp + next + case SENDER + with e i `b1\vbits` show ?thesis using assms by simp + next + case VALUE + with e i `b1\vbits` show ?thesis using assms by simp + next + case TRUE + with e i `b1\vbits` show ?thesis using assms by simp + next + case FALSE + with e i `b1\vbits` show ?thesis using assms by simp + next + case (LVAL x7) + with e i `b1\vbits` show ?thesis using assms by simp + next + case (PLUS x81 x82) + with e i `b1\vbits` show ?thesis using assms by simp + next + case (MINUS x91 x92) + with e i `b1\vbits` show ?thesis using assms by simp + next + case (EQUAL x101 x102) + with e i `b1\vbits` show ?thesis using assms by simp + next + case (LESS x111 x112) + with e i `b1\vbits` show ?thesis using assms by simp + next + case (AND x121 x122) + with e i `b1\vbits` show ?thesis using assms by simp + next + case (OR x131 x132) + with e i `b1\vbits` show ?thesis using assms by simp + next + case (NOT x131) + with e i `b1\vbits` show ?thesis using assms by simp + next + case (CALL x181 x182) + with e i `b1\vbits` show ?thesis using assms by simp + next + case (ECALL x191 x192 x193) + with e i `b1\vbits` show ?thesis using assms by simp + next + case CONTRACTS + with e i `b1\vbits` show ?thesis using assms by simp + qed + next + assume "\ b1\vbits" + with e i show ?thesis using assms by simp + qed + next + case u: (UINT b1 v1) + show ?thesis + proof cases + assume "b1\vbits" + show ?thesis + proof (cases "eupdate e2") + case i: (INT b2 v2) + then show ?thesis + proof cases + assume "b2\vbits" + show ?thesis + proof cases + assume "b1vbits` `b2\vbits` `b1 v1=v2" + with assms show ?thesis using e i u `b1\vbits` `b2\vbits` `b1 b1vbits` show ?thesis using assms by simp + qed + next + assume "b2\vbits" + with e i u `b1\vbits` show ?thesis using assms by simp + qed + next + case u2: (UINT b2 v2) + then show ?thesis + proof cases + assume "b2\vbits" + show ?thesis + proof cases + assume "v1=v2" + with assms show ?thesis using e u u2 `b1\vbits` `b2\vbits` by simp + next + assume "\ v1=v2" + with assms show ?thesis using e u u2 `b1\vbits` `b2\vbits` by simp + qed + next + assume "\b2\vbits" + with e u u2 `b1\vbits` show ?thesis using assms by simp + qed + next + case (ADDRESS x3) + with e u `b1\vbits` show ?thesis using assms by simp + next + case (BALANCE x4) + with e u `b1\vbits` show ?thesis using assms by simp + next + case THIS + with e u `b1\vbits` show ?thesis using assms by simp + next + case SENDER + with e u `b1\vbits` show ?thesis using assms by simp + next + case VALUE + with e u `b1\vbits` show ?thesis using assms by simp + next + case TRUE + with e u `b1\vbits` show ?thesis using assms by simp + next + case FALSE + with e u `b1\vbits` show ?thesis using assms by simp + next + case (LVAL x7) + with e u `b1\vbits` show ?thesis using assms by simp + next + case (PLUS x81 x82) + with e u `b1\vbits` show ?thesis using assms by simp + next + case (MINUS x91 x92) + with e u `b1\vbits` show ?thesis using assms by simp + next + case (EQUAL x101 x102) + with e u `b1\vbits` show ?thesis using assms by simp + next + case (LESS x111 x112) + with e u `b1\vbits` show ?thesis using assms by simp + next + case (AND x121 x122) + with e u `b1\vbits` show ?thesis using assms by simp + next + case (OR x131 x132) + with e u `b1\vbits` show ?thesis using assms by simp + next + case (NOT x131) + with e u `b1\vbits` show ?thesis using assms by simp + next + case (CALL x181 x182) + with e u `b1\vbits` show ?thesis using assms by simp + next + case (ECALL x191 x192 x193) + with e u `b1\vbits` show ?thesis using assms by simp + next + case CONTRACTS + with e u `b1\vbits` show ?thesis using assms by simp + qed + next + assume "\ b1\vbits" + with e u show ?thesis using assms by simp + qed + next + case (ADDRESS x3) + with e show ?thesis using assms by simp + next + case (BALANCE x4) + with e show ?thesis using assms by simp + next + case THIS + with e show ?thesis using assms by simp + next + case SENDER + with e show ?thesis using assms by simp + next + case VALUE + with e show ?thesis using assms by simp + next + case TRUE + with e show ?thesis using assms by simp + next + case FALSE + with e show ?thesis using assms by simp + next + case (LVAL x7) + with e show ?thesis using assms by simp + next + case (PLUS x81 x82) + with e show ?thesis using assms by simp + next + case (MINUS x91 x92) + with e show ?thesis using assms by simp + next + case (EQUAL x101 x102) + with e show ?thesis using assms by simp + next + case (LESS x111 x112) + with e show ?thesis using assms by simp + next + case (AND x121 x122) + with e show ?thesis using assms by simp + next + case (OR x131 x132) + with e show ?thesis using assms by simp + next + case (NOT x131) + with e show ?thesis using assms by simp + next + case (CALL x181 x182) + with e show ?thesis using assms by simp + next + case (ECALL x191 x192 x193) + with e show ?thesis using assms by simp + next + case CONTRACTS + with e show ?thesis using assms by simp + qed +next + case l: (LESS e1 e2) + show ?thesis + proof (cases "eupdate e1") + case i: (INT b1 v1) + show ?thesis + proof cases + assume "b1\vbits" + show ?thesis + proof (cases "eupdate e2") + case i2: (INT b2 v2) + then show ?thesis + proof cases + assume "b2\vbits" + show ?thesis + proof cases + assume "v1vbits` `b2\vbits` by simp + next + assume "\ v1vbits` `b2\vbits` by simp + qed + next + assume "b2\vbits" + with l i i2 `b1\vbits` show ?thesis using assms by simp + qed + next + case u: (UINT b2 v2) + then show ?thesis + proof cases + assume "b2\vbits" + show ?thesis + proof cases + assume "b2vbits` `b2\vbits` `b2 v1vbits` `b2\vbits` `b2 b2vbits` show ?thesis using assms by simp + qed + next + assume "b2\vbits" + with l i u `b1\vbits` show ?thesis using assms by simp + qed + next + case (ADDRESS x3) + with l i `b1\vbits` show ?thesis using assms by simp + next + case (BALANCE x4) + with l i `b1\vbits` show ?thesis using assms by simp + next + case THIS + with l i `b1\vbits` show ?thesis using assms by simp + next + case SENDER + with l i `b1\vbits` show ?thesis using assms by simp + next + case VALUE + with l i `b1\vbits` show ?thesis using assms by simp + next + case TRUE + with l i `b1\vbits` show ?thesis using assms by simp + next + case FALSE + with l i `b1\vbits` show ?thesis using assms by simp + next + case (LVAL x7) + with l i `b1\vbits` show ?thesis using assms by simp + next + case (PLUS x81 x82) + with l i `b1\vbits` show ?thesis using assms by simp + next + case (MINUS x91 x92) + with l i `b1\vbits` show ?thesis using assms by simp + next + case (EQUAL x101 x102) + with l i `b1\vbits` show ?thesis using assms by simp + next + case (LESS x111 x112) + with l i `b1\vbits` show ?thesis using assms by simp + next + case (AND x121 x122) + with l i `b1\vbits` show ?thesis using assms by simp + next + case (OR x131 x132) + with l i `b1\vbits` show ?thesis using assms by simp + next + case (NOT x131) + with l i `b1\vbits` show ?thesis using assms by simp + next + case (CALL x181 x182) + with l i `b1\vbits` show ?thesis using assms by simp + next + case (ECALL x191 x192 x193) + with l i `b1\vbits` show ?thesis using assms by simp + next + case CONTRACTS + with l i `b1\vbits` show ?thesis using assms by simp + qed + next + assume "\ b1\vbits" + with l i show ?thesis using assms by simp + qed + next + case u: (UINT b1 v1) + show ?thesis + proof cases + assume "b1\vbits" + show ?thesis + proof (cases "eupdate e2") + case i: (INT b2 v2) + then show ?thesis + proof cases + assume "b2\vbits" + show ?thesis + proof cases + assume "b1vbits` `b2\vbits` `b1 v1vbits` `b2\vbits` `b1 b1vbits` show ?thesis using assms by simp + qed + next + assume "b2\vbits" + with l i u `b1\vbits` show ?thesis using assms by simp + qed + next + case u2: (UINT b2 v2) + then show ?thesis + proof cases + assume "b2\vbits" + show ?thesis + proof cases + assume "v1vbits` `b2\vbits` by simp + next + assume "\ v1vbits` `b2\vbits` by simp + qed + next + assume "\b2\vbits" + with l u u2 `b1\vbits` show ?thesis using assms by simp + qed + next + case (ADDRESS x3) + with l u `b1\vbits` show ?thesis using assms by simp + next + case (BALANCE x4) + with l u `b1\vbits` show ?thesis using assms by simp + next + case THIS + with l u `b1\vbits` show ?thesis using assms by simp + next + case SENDER + with l u `b1\vbits` show ?thesis using assms by simp + next + case VALUE + with l u `b1\vbits` show ?thesis using assms by simp + next + case TRUE + with l u `b1\vbits` show ?thesis using assms by simp + next + case FALSE + with l u `b1\vbits` show ?thesis using assms by simp + next + case (LVAL x7) + with l u `b1\vbits` show ?thesis using assms by simp + next + case (PLUS x81 x82) + with l u `b1\vbits` show ?thesis using assms by simp + next + case (MINUS x91 x92) + with l u `b1\vbits` show ?thesis using assms by simp + next + case (EQUAL x101 x102) + with l u `b1\vbits` show ?thesis using assms by simp + next + case (LESS x111 x112) + with l u `b1\vbits` show ?thesis using assms by simp + next + case (AND x121 x122) + with l u `b1\vbits` show ?thesis using assms by simp + next + case (OR x131 x132) + with l u `b1\vbits` show ?thesis using assms by simp + next + case (NOT x131) + with l u `b1\vbits` show ?thesis using assms by simp + next + case (CALL x181 x182) + with l u `b1\vbits` show ?thesis using assms by simp + next + case (ECALL x191 x192 x193) + with l u `b1\vbits` show ?thesis using assms by simp + next + case CONTRACTS + with l u `b1\vbits` show ?thesis using assms by simp + qed + next + assume "\ b1\vbits" + with l u show ?thesis using assms by simp + qed + next + case (ADDRESS x3) + with l show ?thesis using assms by simp + next + case (BALANCE x4) + with l show ?thesis using assms by simp + next + case THIS + with l show ?thesis using assms by simp + next + case SENDER + with l show ?thesis using assms by simp + next + case VALUE + with l show ?thesis using assms by simp + next + case TRUE + with l show ?thesis using assms by simp + next + case FALSE + with l show ?thesis using assms by simp + next + case (LVAL x7) + with l show ?thesis using assms by simp + next + case (PLUS x81 x82) + with l show ?thesis using assms by simp + next + case (MINUS x91 x92) + with l show ?thesis using assms by simp + next + case (EQUAL x101 x102) + with l show ?thesis using assms by simp + next + case (LESS x111 x112) + with l show ?thesis using assms by simp + next + case (AND x121 x122) + with l show ?thesis using assms by simp + next + case (OR x131 x132) + with l show ?thesis using assms by simp + next + case (NOT x131) + with l show ?thesis using assms by simp + next + case (CALL x181 x182) + with l show ?thesis using assms by simp + next + case (ECALL x191 x192 x193) + with l show ?thesis using assms by simp + next + case CONTRACTS + with l show ?thesis using assms by simp + qed +next + case a: (AND e1 e2) + show ?thesis + proof (cases "eupdate e1") + case (INT x11 x12) + with a show ?thesis using assms by simp + next + case (UINT x21 x22) + with a show ?thesis using assms by simp + next + case (ADDRESS x3) + with a show ?thesis using assms by simp + next + case (BALANCE x4) + with a show ?thesis using assms by simp + next + case THIS + with a show ?thesis using assms by simp + next + case SENDER + with a show ?thesis using assms by simp + next + case VALUE + with a show ?thesis using assms by simp + next + case t: TRUE + show ?thesis + proof (cases "eupdate e2") + case (INT x11 x12) + with a t show ?thesis using assms by simp + next + case (UINT x21 x22) + with a t show ?thesis using assms by simp + next + case (ADDRESS x3) + with a t show ?thesis using assms by simp + next + case (BALANCE x4) + with a t show ?thesis using assms by simp + next + case THIS + with a t show ?thesis using assms by simp + next + case SENDER + with a t show ?thesis using assms by simp + next + case VALUE + with a t show ?thesis using assms by simp + next + case TRUE + with a t show ?thesis using assms by simp + next + case FALSE + with a t show ?thesis using assms by simp + next + case (LVAL x7) + with a t show ?thesis using assms by simp + next + case (PLUS x81 x82) + with a t show ?thesis using assms by simp + next + case (MINUS x91 x92) + with a t show ?thesis using assms by simp + next + case (EQUAL x101 x102) + with a t show ?thesis using assms by simp + next + case (LESS x111 x112) + with a t show ?thesis using assms by simp + next + case (AND x121 x122) + with a t show ?thesis using assms by simp + next + case (OR x131 x132) + with a t show ?thesis using assms by simp + next + case (NOT x131) + with a t show ?thesis using assms by simp + next + case (CALL x181 x182) + with a t show ?thesis using assms by simp + next + case (ECALL x191 x192 x193) + with a t show ?thesis using assms by simp + next + case CONTRACTS + with a t show ?thesis using assms by simp + qed + next + case f: FALSE + show ?thesis + proof (cases "eupdate e2") + case (INT x11 x12) + with a f show ?thesis using assms by simp + next + case (UINT x21 x22) + with a f show ?thesis using assms by simp + next + case (ADDRESS x3) + with a f show ?thesis using assms by simp + next + case (BALANCE x4) + with a f show ?thesis using assms by simp + next + case THIS + with a f show ?thesis using assms by simp + next + case SENDER + with a f show ?thesis using assms by simp + next + case VALUE + with a f show ?thesis using assms by simp + next + case TRUE + with a f show ?thesis using assms by simp + next + case FALSE + with a f show ?thesis using assms by simp + next + case (LVAL x7) + with a f show ?thesis using assms by simp + next + case (PLUS x81 x82) + with a f show ?thesis using assms by simp + next + case (MINUS x91 x92) + with a f show ?thesis using assms by simp + next + case (EQUAL x101 x102) + with a f show ?thesis using assms by simp + next + case (LESS x111 x112) + with a f show ?thesis using assms by simp + next + case (AND x121 x122) + with a f show ?thesis using assms by simp + next + case (OR x131 x132) + with a f show ?thesis using assms by simp + next + case (NOT x131) + with a f show ?thesis using assms by simp + next + case (CALL x181 x182) + with a f show ?thesis using assms by simp + next + case (ECALL x191 x192 x193) + with a f show ?thesis using assms by simp + next + case CONTRACTS + with a f show ?thesis using assms by simp + qed + next + case (LVAL x7) + with a show ?thesis using assms by simp + next + case (PLUS x81 x82) + with a show ?thesis using assms by simp + next + case (MINUS x91 x92) + with a show ?thesis using assms by simp + next + case (EQUAL x101 x102) + with a show ?thesis using assms by simp + next + case (LESS x111 x112) + with a show ?thesis using assms by simp + next + case (AND x121 x122) + with a show ?thesis using assms by simp + next + case (OR x131 x132) + with a show ?thesis using assms by simp + next + case (NOT x131) + with a show ?thesis using assms by simp + next + case (CALL x181 x182) + with a show ?thesis using assms by simp + next + case (ECALL x191 x192 x193) + with a show ?thesis using assms by simp + next + case CONTRACTS + with a show ?thesis using assms by simp + qed +next + case o: (OR e1 e2) + show ?thesis + proof (cases "eupdate e1") + case (INT x11 x12) + with o show ?thesis using assms by simp + next + case (UINT x21 x22) + with o show ?thesis using assms by simp + next + case (ADDRESS x3) + with o show ?thesis using assms by simp + next + case (BALANCE x4) + with o show ?thesis using assms by simp + next + case THIS + with o show ?thesis using assms by simp + next + case SENDER + with o show ?thesis using assms by simp + next + case VALUE + with o show ?thesis using assms by simp + next + case t: TRUE + show ?thesis + proof (cases "eupdate e2") + case (INT x11 x12) + with o t show ?thesis using assms by simp + next + case (UINT x21 x22) + with o t show ?thesis using assms by simp + next + case (ADDRESS x3) + with o t show ?thesis using assms by simp + next + case (BALANCE x4) + with o t show ?thesis using assms by simp + next + case THIS + with o t show ?thesis using assms by simp + next + case SENDER + with o t show ?thesis using assms by simp + next + case VALUE + with o t show ?thesis using assms by simp + next + case TRUE + with o t show ?thesis using assms by simp + next + case FALSE + with o t show ?thesis using assms by simp + next + case (LVAL x7) + with o t show ?thesis using assms by simp + next + case (PLUS x81 x82) + with o t show ?thesis using assms by simp + next + case (MINUS x91 x92) + with o t show ?thesis using assms by simp + next + case (EQUAL x101 x102) + with o t show ?thesis using assms by simp + next + case (LESS x111 x112) + with o t show ?thesis using assms by simp + next + case (AND x121 x122) + with o t show ?thesis using assms by simp + next + case (OR x131 x132) + with o t show ?thesis using assms by simp + next + case (NOT x131) + with o t show ?thesis using assms by simp + next + case (CALL x181 x182) + with o t show ?thesis using assms by simp + next + case (ECALL x191 x192 x193) + with o t show ?thesis using assms by simp + next + case CONTRACTS + with o t show ?thesis using assms by simp + qed + next + case f: FALSE + show ?thesis + proof (cases "eupdate e2") + case (INT x11 x12) + with o f show ?thesis using assms by simp + next + case (UINT x21 x22) + with o f show ?thesis using assms by simp + next + case (ADDRESS x3) + with o f show ?thesis using assms by simp + next + case (BALANCE x4) + with o f show ?thesis using assms by simp + next + case THIS + with o f show ?thesis using assms by simp + next + case SENDER + with o f show ?thesis using assms by simp + next + case VALUE + with o f show ?thesis using assms by simp + next + case TRUE + with o f show ?thesis using assms by simp + next + case FALSE + with o f show ?thesis using assms by simp + next + case (LVAL x7) + with o f show ?thesis using assms by simp + next + case (PLUS x81 x82) + with o f show ?thesis using assms by simp + next + case (MINUS x91 x92) + with o f show ?thesis using assms by simp + next + case (EQUAL x101 x102) + with o f show ?thesis using assms by simp + next + case (LESS x111 x112) + with o f show ?thesis using assms by simp + next + case (AND x121 x122) + with o f show ?thesis using assms by simp + next + case (OR x131 x132) + with o f show ?thesis using assms by simp + next + case (NOT x131) + with o f show ?thesis using assms by simp + next + case (CALL x181 x182) + with o f show ?thesis using assms by simp + next + case (ECALL x191 x192 x193) + with o f show ?thesis using assms by simp + next + case CONTRACTS + with o f show ?thesis using assms by simp + qed + next + case (LVAL x7) + with o show ?thesis using assms by simp + next + case (PLUS x81 x82) + with o show ?thesis using assms by simp + next + case (MINUS x91 x92) + with o show ?thesis using assms by simp + next + case (EQUAL x101 x102) + with o show ?thesis using assms by simp + next + case (LESS x111 x112) + with o show ?thesis using assms by simp + next + case (AND x121 x122) + with o show ?thesis using assms by simp + next + case (OR x131 x132) + with o show ?thesis using assms by simp + next + case (NOT x131) + with o show ?thesis using assms by simp + next + case (CALL x181 x182) + with o show ?thesis using assms by simp + next + case (ECALL x191 x192 x193) + with o show ?thesis using assms by simp + next + case CONTRACTS + with o show ?thesis using assms by simp + qed +next + case o: (NOT e1) + show ?thesis + proof (cases "eupdate e1") + case (INT x11 x12) + with o show ?thesis using assms by simp + next + case (UINT x21 x22) + with o show ?thesis using assms by simp + next + case (ADDRESS x3) + with o show ?thesis using assms by simp + next + case (BALANCE x4) + with o show ?thesis using assms by simp + next + case THIS + with o show ?thesis using assms by simp + next + case SENDER + with o show ?thesis using assms by simp + next + case VALUE + with o show ?thesis using assms by simp + next + case t: TRUE + with o show ?thesis using assms by simp + next + case f: FALSE + with o show ?thesis using assms by simp + next + case (LVAL x7) + with o show ?thesis using assms by simp + next + case (PLUS x81 x82) + with o show ?thesis using assms by simp + next + case (MINUS x91 x92) + with o show ?thesis using assms by simp + next + case (EQUAL x101 x102) + with o show ?thesis using assms by simp + next + case (LESS x111 x112) + with o show ?thesis using assms by simp + next + case (AND x121 x122) + with o show ?thesis using assms by simp + next + case (OR x131 x132) + with o show ?thesis using assms by simp + next + case (NOT x131) + with o show ?thesis using assms by simp + next + case (CALL x181 x182) + with o show ?thesis using assms by simp + next + case (ECALL x191 x192 x193) + with o show ?thesis using assms by simp + next + case CONTRACTS + with o show ?thesis using assms by simp + qed +next + case (CALL x181 x182) + with assms show ?thesis by simp +next + case (ECALL x191 x192 x193) + with assms show ?thesis by simp +next + case CONTRACTS + with assms show ?thesis by simp +qed + +lemma update_bounds_uint: + assumes "eupdate ex = UINT b v" and "b\vbits" + shows "v < 2^b \ v \ 0" +proof (cases ex) + case (INT b' v') + with assms show ?thesis + proof cases + assume "b'\vbits" + show ?thesis + proof cases + assume "v'\0" + with INT show ?thesis using assms `b'\vbits` by simp + next + assume "\ v'\0" + with INT show ?thesis using assms `b'\vbits` by simp + qed + next + assume "\ b'\vbits" + with INT show ?thesis using assms by simp + qed +next + case (UINT b' v') + then show ?thesis + proof cases + assume "b'\vbits" + with UINT show ?thesis using assms by auto + next + assume "\ b'\vbits" + with UINT show ?thesis using assms by auto + qed +next + case (ADDRESS x3) + with assms show ?thesis by simp +next + case (BALANCE x4) + with assms show ?thesis by simp +next + case THIS + with assms show ?thesis by simp +next + case SENDER + with assms show ?thesis by simp +next + case VALUE + with assms show ?thesis by simp +next + case TRUE + with assms show ?thesis by simp +next + case FALSE + with assms show ?thesis by simp +next + case (LVAL x7) + with assms show ?thesis by simp +next + case p: (PLUS e1 e2) + show ?thesis + proof (cases "eupdate e1") + case i: (INT b1 v1) + with p show ?thesis + proof cases + assume "b1\vbits" + show ?thesis + proof (cases "eupdate e2") + case i2: (INT b2 v2) + then show ?thesis + proof cases + let ?v="v1+v2" + assume "b2\vbits" + show ?thesis + proof cases + assume "?v\0" + with assms show ?thesis using p i i2 `b1\vbits` `b2\vbits` by simp + next + assume "\?v\0" + with assms show ?thesis using p i i2 `b1\vbits` `b2\vbits` by simp + qed + next + assume "b2\vbits" + with p i i2 `b1\vbits` show ?thesis using assms by simp + qed + next + case u: (UINT b2 v2) + then show ?thesis + proof cases + let ?v="v1+v2" + assume "b2\vbits" + show ?thesis + proof cases + assume "b20" + with assms show ?thesis using p i u `b1\vbits` `b2\vbits` `b2?v\0" + with assms show ?thesis using p i u `b1\vbits` `b2\vbits` `b2 b2vbits` show ?thesis using assms by simp + qed + next + assume "b2\vbits" + with p i u `b1\vbits` show ?thesis using assms by simp + qed + next + case (ADDRESS x3) + with p i `b1\vbits` show ?thesis using assms by simp + next + case (BALANCE x4) + with p i `b1\vbits` show ?thesis using assms by simp + next + case THIS + with p i `b1\vbits` show ?thesis using assms by simp + next + case SENDER + with p i `b1\vbits` show ?thesis using assms by simp + next + case VALUE + with p i `b1\vbits` show ?thesis using assms by simp + next + case TRUE + with p i `b1\vbits` show ?thesis using assms by simp + next + case FALSE + with p i `b1\vbits` show ?thesis using assms by simp + next + case (LVAL x7) + with p i `b1\vbits` show ?thesis using assms by simp + next + case (PLUS x81 x82) + with p i `b1\vbits` show ?thesis using assms by simp + next + case (MINUS x91 x92) + with p i `b1\vbits` show ?thesis using assms by simp + next + case (EQUAL x101 x102) + with p i `b1\vbits` show ?thesis using assms by simp + next + case (LESS x111 x112) + with p i `b1\vbits` show ?thesis using assms by simp + next + case (AND x121 x122) + with p i `b1\vbits` show ?thesis using assms by simp + next + case (OR x131 x132) + with p i `b1\vbits` show ?thesis using assms by simp + next + case (NOT x131) + with p i `b1\vbits` show ?thesis using assms by simp + next + case (CALL x181 x182) + with p i `b1\vbits` show ?thesis using assms by simp + next + case (ECALL x191 x192 x193) + with p i `b1\vbits` show ?thesis using assms by simp + next + case CONTRACTS + with p i `b1\vbits` show ?thesis using assms by simp + qed + next + assume "\ b1\vbits" + with p i show ?thesis using assms by simp + qed + next + case u: (UINT b1 v1) + with p show ?thesis + proof cases + assume "b1\vbits" + show ?thesis + proof (cases "eupdate e2") + case i: (INT b2 v2) + then show ?thesis + proof cases + let ?v="v1+v2" + assume "b2\vbits" + show ?thesis + proof cases + assume "b10" + with assms show ?thesis using p i u `b1\vbits` `b2\vbits` `b1?v\0" + with assms show ?thesis using p i u `b1\vbits` `b2\vbits` `b1 b1vbits` show ?thesis using assms by simp + qed + next + assume "b2\vbits" + with p i u `b1\vbits` show ?thesis using assms by simp + qed + next + case u2: (UINT b2 v2) + then show ?thesis + proof cases + let ?x="((v1 + v2) mod (2^(max b1 b2)))" + assume "b2\vbits" + with `b1\vbits` u u2 have "eupdate (PLUS e1 e2) = UINT (max b1 b2) ?x" by simp + with assms have "b=max b1 b2" and "v=?x" using p by (simp,simp) + moreover from `b1\vbits` have "max b1 b2>0" by auto + hence "?x < 2 ^(max b1 b2)" by simp + moreover have "?x \ 0" by simp + ultimately show ?thesis by simp + next + assume "\b2\vbits" + with p u u2 `b1\vbits` show ?thesis using assms by simp + qed + next + case (ADDRESS x3) + with p u `b1\vbits` show ?thesis using assms by simp + next + case (BALANCE x4) + with p u `b1\vbits` show ?thesis using assms by simp + next + case THIS + with p u `b1\vbits` show ?thesis using assms by simp + next + case SENDER + with p u `b1\vbits` show ?thesis using assms by simp + next + case VALUE + with p u `b1\vbits` show ?thesis using assms by simp + next + case TRUE + with p u `b1\vbits` show ?thesis using assms by simp + next + case FALSE + with p u `b1\vbits` show ?thesis using assms by simp + next + case (LVAL x7) + with p u `b1\vbits` show ?thesis using assms by simp + next + case (PLUS x81 x82) + with p u `b1\vbits` show ?thesis using assms by simp + next + case (MINUS x91 x92) + with p u `b1\vbits` show ?thesis using assms by simp + next + case (EQUAL x101 x102) + with p u `b1\vbits` show ?thesis using assms by simp + next + case (LESS x111 x112) + with p u `b1\vbits` show ?thesis using assms by simp + next + case (AND x121 x122) + with p u `b1\vbits` show ?thesis using assms by simp + next + case (OR x131 x132) + with p u `b1\vbits` show ?thesis using assms by simp + next + case (NOT x131) + with p u `b1\vbits` show ?thesis using assms by simp + next + case (CALL x181 x182) + with p u `b1\vbits` show ?thesis using assms by simp + next + case (ECALL x191 x192 x193) + with p u `b1\vbits` show ?thesis using assms by simp + next + case CONTRACTS + with p u `b1\vbits` show ?thesis using assms by simp + qed + next + assume "\ b1\vbits" + with p u show ?thesis using assms by simp + qed + next + case (ADDRESS x3) + with p show ?thesis using assms by simp + next + case (BALANCE x4) + with p show ?thesis using assms by simp + next + case THIS + with p show ?thesis using assms by simp + next + case SENDER + with p show ?thesis using assms by simp + next + case VALUE + with p show ?thesis using assms by simp + next + case TRUE + with p show ?thesis using assms by simp + next + case FALSE + with p show ?thesis using assms by simp + next + case (LVAL x7) + with p show ?thesis using assms by simp + next + case (PLUS x81 x82) + with p show ?thesis using assms by simp + next + case (MINUS x91 x92) + with p show ?thesis using assms by simp + next + case (EQUAL x101 x102) + with p show ?thesis using assms by simp + next + case (LESS x111 x112) + with p show ?thesis using assms by simp + next + case (AND x121 x122) + with p show ?thesis using assms by simp + next + case (OR x131 x132) + with p show ?thesis using assms by simp + next + case (NOT x131) + with p show ?thesis using assms by simp + next + case (CALL x181 x182) + with p show ?thesis using assms by simp + next + case (ECALL x191 x192 x193) + with p show ?thesis using assms by simp + next + case CONTRACTS + with p show ?thesis using assms by simp + qed +next + case m: (MINUS e1 e2) + show ?thesis + proof (cases "eupdate e1") + case i: (INT b1 v1) + with m show ?thesis + proof cases + assume "b1\vbits" + show ?thesis + proof (cases "eupdate e2") + case i2: (INT b2 v2) + then show ?thesis + proof cases + let ?v="v1-v2" + assume "b2\vbits" + show ?thesis + proof cases + assume "?v\0" + with assms show ?thesis using m i i2 `b1\vbits` `b2\vbits` by simp + next + assume "\?v\0" + with assms show ?thesis using m i i2 `b1\vbits` `b2\vbits` by simp + qed + next + assume "b2\vbits" + with m i i2 `b1\vbits` show ?thesis using assms by simp + qed + next + case u: (UINT b2 v2) + then show ?thesis + proof cases + let ?v="v1-v2" + assume "b2\vbits" + show ?thesis + proof cases + assume "b20" + with assms show ?thesis using m i u `b1\vbits` `b2\vbits` `b2?v\0" + with assms show ?thesis using m i u `b1\vbits` `b2\vbits` `b2 b2vbits` show ?thesis using assms by simp + qed + next + assume "b2\vbits" + with m i u `b1\vbits` show ?thesis using assms by simp + qed + next + case (ADDRESS x3) + with m i `b1\vbits` show ?thesis using assms by simp + next + case (BALANCE x4) + with m i `b1\vbits` show ?thesis using assms by simp + next + case THIS + with m i `b1\vbits` show ?thesis using assms by simp + next + case SENDER + with m i `b1\vbits` show ?thesis using assms by simp + next + case VALUE + with m i `b1\vbits` show ?thesis using assms by simp + next + case TRUE + with m i `b1\vbits` show ?thesis using assms by simp + next + case FALSE + with m i `b1\vbits` show ?thesis using assms by simp + next + case (LVAL x7) + with m i `b1\vbits` show ?thesis using assms by simp + next + case (PLUS x81 x82) + with m i `b1\vbits` show ?thesis using assms by simp + next + case (MINUS x91 x92) + with m i `b1\vbits` show ?thesis using assms by simp + next + case (EQUAL x101 x102) + with m i `b1\vbits` show ?thesis using assms by simp + next + case (LESS x111 x112) + with m i `b1\vbits` show ?thesis using assms by simp + next + case (AND x121 x122) + with m i `b1\vbits` show ?thesis using assms by simp + next + case (OR x131 x132) + with m i `b1\vbits` show ?thesis using assms by simp + next + case (NOT x131) + with m i `b1\vbits` show ?thesis using assms by simp + next + case (CALL x181 x182) + with m i `b1\vbits` show ?thesis using assms by simp + next + case (ECALL x191 x192 x193) + with m i `b1\vbits` show ?thesis using assms by simp + next + case CONTRACTS + with m i `b1\vbits` show ?thesis using assms by simp + qed + next + assume "\ b1\vbits" + with m i show ?thesis using assms by simp + qed + next + case u: (UINT b1 v1) + with m show ?thesis + proof cases + assume "b1\vbits" + show ?thesis + proof (cases "eupdate e2") + case i: (INT b2 v2) + then show ?thesis + proof cases + let ?v="v1-v2" + assume "b2\vbits" + show ?thesis + proof cases + assume "b10" + with assms show ?thesis using m i u `b1\vbits` `b2\vbits` `b1?v\0" + with assms show ?thesis using m i u `b1\vbits` `b2\vbits` `b1 b1vbits` show ?thesis using assms by simp + qed + next + assume "b2\vbits" + with m i u `b1\vbits` show ?thesis using assms by simp + qed + next + case u2: (UINT b2 v2) + then show ?thesis + proof cases + let ?x="((v1 - v2) mod (2^(max b1 b2)))" + assume "b2\vbits" + with `b1\vbits` u u2 have "eupdate (MINUS e1 e2) = UINT (max b1 b2) ?x" by simp + with assms have "b=max b1 b2" and "v=?x" using m by (simp,simp) + moreover from `b1\vbits` have "max b1 b2>0" by auto + hence "?x < 2 ^(max b1 b2)" by simp + moreover have "?x \ 0" by simp + ultimately show ?thesis by simp + next + assume "\b2\vbits" + with m u u2 `b1\vbits` show ?thesis using assms by simp + qed + next + case (ADDRESS x3) + with m u `b1\vbits` show ?thesis using assms by simp + next + case (BALANCE x4) + with m u `b1\vbits` show ?thesis using assms by simp + next + case THIS + with m u `b1\vbits` show ?thesis using assms by simp + next + case SENDER + with m u `b1\vbits` show ?thesis using assms by simp + next + case VALUE + with m u `b1\vbits` show ?thesis using assms by simp + next + case TRUE + with m u `b1\vbits` show ?thesis using assms by simp + next + case FALSE + with m u `b1\vbits` show ?thesis using assms by simp + next + case (LVAL x7) + with m u `b1\vbits` show ?thesis using assms by simp + next + case (PLUS x81 x82) + with m u `b1\vbits` show ?thesis using assms by simp + next + case (MINUS x91 x92) + with m u `b1\vbits` show ?thesis using assms by simp + next + case (EQUAL x101 x102) + with m u `b1\vbits` show ?thesis using assms by simp + next + case (LESS x111 x112) + with m u `b1\vbits` show ?thesis using assms by simp + next + case (AND x121 x122) + with m u `b1\vbits` show ?thesis using assms by simp + next + case (OR x131 x132) + with m u `b1\vbits` show ?thesis using assms by simp + next + case (NOT x131) + with m u `b1\vbits` show ?thesis using assms by simp + next + case (CALL x181 x182) + with m u `b1\vbits` show ?thesis using assms by simp + next + case (ECALL x191 x192 x193) + with m u `b1\vbits` show ?thesis using assms by simp + next + case CONTRACTS + with m u `b1\vbits` show ?thesis using assms by simp + qed + next + assume "\ b1\vbits" + with m u show ?thesis using assms by simp + qed + next + case (ADDRESS x3) + with m show ?thesis using assms by simp + next + case (BALANCE x4) + with m show ?thesis using assms by simp + next + case THIS + with m show ?thesis using assms by simp + next + case SENDER + with m show ?thesis using assms by simp + next + case VALUE + with m show ?thesis using assms by simp + next + case TRUE + with m show ?thesis using assms by simp + next + case FALSE + with m show ?thesis using assms by simp + next + case (LVAL x7) + with m show ?thesis using assms by simp + next + case (PLUS x81 x82) + with m show ?thesis using assms by simp + next + case (MINUS x91 x92) + with m show ?thesis using assms by simp + next + case (EQUAL x101 x102) + with m show ?thesis using assms by simp + next + case (LESS x111 x112) + with m show ?thesis using assms by simp + next + case (AND x121 x122) + with m show ?thesis using assms by simp + next + case (OR x131 x132) + with m show ?thesis using assms by simp + next + case (NOT x131) + with m show ?thesis using assms by simp + next + case (CALL x181 x182) + with m show ?thesis using assms by simp + next + case (ECALL x191 x192 x193) + with m show ?thesis using assms by simp + next + case CONTRACTS + with m show ?thesis using assms by simp + qed +next + case e: (EQUAL e1 e2) + show ?thesis + proof (cases "eupdate e1") + case i: (INT b1 v1) + show ?thesis + proof cases + assume "b1\vbits" + show ?thesis + proof (cases "eupdate e2") + case i2: (INT b2 v2) + then show ?thesis + proof cases + assume "b2\vbits" + show ?thesis + proof cases + assume "v1=v2" + with assms show ?thesis using e i i2 `b1\vbits` `b2\vbits` by simp + next + assume "\ v1=v2" + with assms show ?thesis using e i i2 `b1\vbits` `b2\vbits` by simp + qed + next + assume "b2\vbits" + with e i i2 `b1\vbits` show ?thesis using assms by simp + qed + next + case u: (UINT b2 v2) + then show ?thesis + proof cases + assume "b2\vbits" + show ?thesis + proof cases + assume "b2vbits` `b2\vbits` `b2 v1=v2" + with assms show ?thesis using e i u `b1\vbits` `b2\vbits` `b2 b2vbits` show ?thesis using assms by simp + qed + next + assume "b2\vbits" + with e i u `b1\vbits` show ?thesis using assms by simp + qed + next + case (ADDRESS x3) + with e i `b1\vbits` show ?thesis using assms by simp + next + case (BALANCE x4) + with e i `b1\vbits` show ?thesis using assms by simp + next + case THIS + with e i `b1\vbits` show ?thesis using assms by simp + next + case SENDER + with e i `b1\vbits` show ?thesis using assms by simp + next + case VALUE + with e i `b1\vbits` show ?thesis using assms by simp + next + case TRUE + with e i `b1\vbits` show ?thesis using assms by simp + next + case FALSE + with e i `b1\vbits` show ?thesis using assms by simp + next + case (LVAL x7) + with e i `b1\vbits` show ?thesis using assms by simp + next + case (PLUS x81 x82) + with e i `b1\vbits` show ?thesis using assms by simp + next + case (MINUS x91 x92) + with e i `b1\vbits` show ?thesis using assms by simp + next + case (EQUAL x101 x102) + with e i `b1\vbits` show ?thesis using assms by simp + next + case (LESS x111 x112) + with e i `b1\vbits` show ?thesis using assms by simp + next + case (AND x121 x122) + with e i `b1\vbits` show ?thesis using assms by simp + next + case (OR x131 x132) + with e i `b1\vbits` show ?thesis using assms by simp + next + case (NOT x131) + with e i `b1\vbits` show ?thesis using assms by simp + next + case (CALL x181 x182) + with e i `b1\vbits` show ?thesis using assms by simp + next + case (ECALL x191 x192 x193) + with e i `b1\vbits` show ?thesis using assms by simp + next + case CONTRACTS + with e i `b1\vbits` show ?thesis using assms by simp + qed + next + assume "\ b1\vbits" + with e i show ?thesis using assms by simp + qed + next + case u: (UINT b1 v1) + show ?thesis + proof cases + assume "b1\vbits" + show ?thesis + proof (cases "eupdate e2") + case i: (INT b2 v2) + then show ?thesis + proof cases + let ?v="v1+v2" + assume "b2\vbits" + show ?thesis + proof cases + assume "b1vbits` `b2\vbits` `b1 v1=v2" + with assms show ?thesis using e i u `b1\vbits` `b2\vbits` `b1 b1vbits` show ?thesis using assms by simp + qed + next + assume "b2\vbits" + with e i u `b1\vbits` show ?thesis using assms by simp + qed + next + case u2: (UINT b2 v2) + then show ?thesis + proof cases + assume "b2\vbits" + show ?thesis + proof cases + assume "v1=v2" + with assms show ?thesis using e u u2 `b1\vbits` `b2\vbits` by simp + next + assume "\ v1=v2" + with assms show ?thesis using e u u2 `b1\vbits` `b2\vbits` by simp + qed + next + assume "\b2\vbits" + with e u u2 `b1\vbits` show ?thesis using assms by simp + qed + next + case (ADDRESS x3) + with e u `b1\vbits` show ?thesis using assms by simp + next + case (BALANCE x4) + with e u `b1\vbits` show ?thesis using assms by simp + next + case THIS + with e u `b1\vbits` show ?thesis using assms by simp + next + case SENDER + with e u `b1\vbits` show ?thesis using assms by simp + next + case VALUE + with e u `b1\vbits` show ?thesis using assms by simp + next + case TRUE + with e u `b1\vbits` show ?thesis using assms by simp + next + case FALSE + with e u `b1\vbits` show ?thesis using assms by simp + next + case (LVAL x7) + with e u `b1\vbits` show ?thesis using assms by simp + next + case (PLUS x81 x82) + with e u `b1\vbits` show ?thesis using assms by simp + next + case (MINUS x91 x92) + with e u `b1\vbits` show ?thesis using assms by simp + next + case (EQUAL x101 x102) + with e u `b1\vbits` show ?thesis using assms by simp + next + case (LESS x111 x112) + with e u `b1\vbits` show ?thesis using assms by simp + next + case (AND x121 x122) + with e u `b1\vbits` show ?thesis using assms by simp + next + case (OR x131 x132) + with e u `b1\vbits` show ?thesis using assms by simp + next + case (NOT x131) + with e u `b1\vbits` show ?thesis using assms by simp + next + case (CALL x181 x182) + with e u `b1\vbits` show ?thesis using assms by simp + next + case (ECALL x191 x192 x193) + with e u `b1\vbits` show ?thesis using assms by simp + next + case CONTRACTS + with e u `b1\vbits` show ?thesis using assms by simp + qed + next + assume "\ b1\vbits" + with e u show ?thesis using assms by simp + qed + next + case (ADDRESS x3) + with e show ?thesis using assms by simp + next + case (BALANCE x4) + with e show ?thesis using assms by simp + next + case THIS + with e show ?thesis using assms by simp + next + case SENDER + with e show ?thesis using assms by simp + next + case VALUE + with e show ?thesis using assms by simp + next + case TRUE + with e show ?thesis using assms by simp + next + case FALSE + with e show ?thesis using assms by simp + next + case (LVAL x7) + with e show ?thesis using assms by simp + next + case (PLUS x81 x82) + with e show ?thesis using assms by simp + next + case (MINUS x91 x92) + with e show ?thesis using assms by simp + next + case (EQUAL x101 x102) + with e show ?thesis using assms by simp + next + case (LESS x111 x112) + with e show ?thesis using assms by simp + next + case (AND x121 x122) + with e show ?thesis using assms by simp + next + case (OR x131 x132) + with e show ?thesis using assms by simp + next + case (NOT x131) + with e show ?thesis using assms by simp + next + case (CALL x181 x182) + with e show ?thesis using assms by simp + next + case (ECALL x191 x192 x193) + with e show ?thesis using assms by simp + next + case CONTRACTS + with e show ?thesis using assms by simp + qed +next + case l: (LESS e1 e2) + show ?thesis + proof (cases "eupdate e1") + case i: (INT b1 v1) + show ?thesis + proof cases + assume "b1\vbits" + show ?thesis + proof (cases "eupdate e2") + case i2: (INT b2 v2) + then show ?thesis + proof cases + assume "b2\vbits" + show ?thesis + proof cases + assume "v1vbits` `b2\vbits` by simp + next + assume "\ v1vbits` `b2\vbits` by simp + qed + next + assume "b2\vbits" + with l i i2 `b1\vbits` show ?thesis using assms by simp + qed + next + case u: (UINT b2 v2) + then show ?thesis + proof cases + assume "b2\vbits" + show ?thesis + proof cases + assume "b2vbits` `b2\vbits` `b2 v1vbits` `b2\vbits` `b2 b2vbits` show ?thesis using assms by simp + qed + next + assume "b2\vbits" + with l i u `b1\vbits` show ?thesis using assms by simp + qed + next + case (ADDRESS x3) + with l i `b1\vbits` show ?thesis using assms by simp + next + case (BALANCE x4) + with l i `b1\vbits` show ?thesis using assms by simp + next + case THIS + with l i `b1\vbits` show ?thesis using assms by simp + next + case SENDER + with l i `b1\vbits` show ?thesis using assms by simp + next + case VALUE + with l i `b1\vbits` show ?thesis using assms by simp + next + case TRUE + with l i `b1\vbits` show ?thesis using assms by simp + next + case FALSE + with l i `b1\vbits` show ?thesis using assms by simp + next + case (LVAL x7) + with l i `b1\vbits` show ?thesis using assms by simp + next + case (PLUS x81 x82) + with l i `b1\vbits` show ?thesis using assms by simp + next + case (MINUS x91 x92) + with l i `b1\vbits` show ?thesis using assms by simp + next + case (EQUAL x101 x102) + with l i `b1\vbits` show ?thesis using assms by simp + next + case (LESS x111 x112) + with l i `b1\vbits` show ?thesis using assms by simp + next + case (AND x121 x122) + with l i `b1\vbits` show ?thesis using assms by simp + next + case (OR x131 x132) + with l i `b1\vbits` show ?thesis using assms by simp + next + case (NOT x131) + with l i `b1\vbits` show ?thesis using assms by simp + next + case (CALL x181 x182) + with l i `b1\vbits` show ?thesis using assms by simp + next + case (ECALL x191 x192 x193) + with l i `b1\vbits` show ?thesis using assms by simp + next + case CONTRACTS + with l i `b1\vbits` show ?thesis using assms by simp + qed + next + assume "\ b1\vbits" + with l i show ?thesis using assms by simp + qed + next + case u: (UINT b1 v1) + show ?thesis + proof cases + assume "b1\vbits" + show ?thesis + proof (cases "eupdate e2") + case i: (INT b2 v2) + then show ?thesis + proof cases + let ?v="v1+v2" + assume "b2\vbits" + show ?thesis + proof cases + assume "b1vbits` `b2\vbits` `b1 v1vbits` `b2\vbits` `b1 b1vbits` show ?thesis using assms by simp + qed + next + assume "b2\vbits" + with l i u `b1\vbits` show ?thesis using assms by simp + qed + next + case u2: (UINT b2 v2) + then show ?thesis + proof cases + assume "b2\vbits" + show ?thesis + proof cases + assume "v1vbits` `b2\vbits` by simp + next + assume "\ v1vbits` `b2\vbits` by simp + qed + next + assume "\b2\vbits" + with l u u2 `b1\vbits` show ?thesis using assms by simp + qed + next + case (ADDRESS x3) + with l u `b1\vbits` show ?thesis using assms by simp + next + case (BALANCE x4) + with l u `b1\vbits` show ?thesis using assms by simp + next + case THIS + with l u `b1\vbits` show ?thesis using assms by simp + next + case SENDER + with l u `b1\vbits` show ?thesis using assms by simp + next + case VALUE + with l u `b1\vbits` show ?thesis using assms by simp + next + case TRUE + with l u `b1\vbits` show ?thesis using assms by simp + next + case FALSE + with l u `b1\vbits` show ?thesis using assms by simp + next + case (LVAL x7) + with l u `b1\vbits` show ?thesis using assms by simp + next + case (PLUS x81 x82) + with l u `b1\vbits` show ?thesis using assms by simp + next + case (MINUS x91 x92) + with l u `b1\vbits` show ?thesis using assms by simp + next + case (EQUAL x101 x102) + with l u `b1\vbits` show ?thesis using assms by simp + next + case (LESS x111 x112) + with l u `b1\vbits` show ?thesis using assms by simp + next + case (AND x121 x122) + with l u `b1\vbits` show ?thesis using assms by simp + next + case (OR x131 x132) + with l u `b1\vbits` show ?thesis using assms by simp + next + case (NOT x131) + with l u `b1\vbits` show ?thesis using assms by simp + next + case (CALL x181 x182) + with l u `b1\vbits` show ?thesis using assms by simp + next + case (ECALL x191 x192 x193) + with l u `b1\vbits` show ?thesis using assms by simp + next + case CONTRACTS + with l u `b1\vbits` show ?thesis using assms by simp + qed + next + assume "\ b1\vbits" + with l u show ?thesis using assms by simp + qed + next + case (ADDRESS x3) + with l show ?thesis using assms by simp + next + case (BALANCE x4) + with l show ?thesis using assms by simp + next + case THIS + with l show ?thesis using assms by simp + next + case SENDER + with l show ?thesis using assms by simp + next + case VALUE + with l show ?thesis using assms by simp + next + case TRUE + with l show ?thesis using assms by simp + next + case FALSE + with l show ?thesis using assms by simp + next + case (LVAL x7) + with l show ?thesis using assms by simp + next + case (PLUS x81 x82) + with l show ?thesis using assms by simp + next + case (MINUS x91 x92) + with l show ?thesis using assms by simp + next + case (EQUAL x101 x102) + with l show ?thesis using assms by simp + next + case (LESS x111 x112) + with l show ?thesis using assms by simp + next + case (AND x121 x122) + with l show ?thesis using assms by simp + next + case (OR x131 x132) + with l show ?thesis using assms by simp + next + case (NOT x131) + with l show ?thesis using assms by simp + next + case (CALL x181 x182) + with l show ?thesis using assms by simp + next + case (ECALL x191 x192 x193) + with l show ?thesis using assms by simp + next + case CONTRACTS + with l show ?thesis using assms by simp + qed +next + case a: (AND e1 e2) + show ?thesis + proof (cases "eupdate e1") + case (INT x11 x12) + with a show ?thesis using assms by simp + next + case (UINT x21 x22) + with a show ?thesis using assms by simp + next + case (ADDRESS x3) + with a show ?thesis using assms by simp + next + case (BALANCE x4) + with a show ?thesis using assms by simp + next + case THIS + with a show ?thesis using assms by simp + next + case SENDER + with a show ?thesis using assms by simp + next + case VALUE + with a show ?thesis using assms by simp + next + case t: TRUE + show ?thesis + proof (cases "eupdate e2") + case (INT x11 x12) + with a t show ?thesis using assms by simp + next + case (UINT x21 x22) + with a t show ?thesis using assms by simp + next + case (ADDRESS x3) + with a t show ?thesis using assms by simp + next + case (BALANCE x4) + with a t show ?thesis using assms by simp + next + case THIS + with a t show ?thesis using assms by simp + next + case SENDER + with a t show ?thesis using assms by simp + next + case VALUE + with a t show ?thesis using assms by simp + next + case TRUE + with a t show ?thesis using assms by simp + next + case FALSE + with a t show ?thesis using assms by simp + next + case (LVAL x7) + with a t show ?thesis using assms by simp + next + case (PLUS x81 x82) + with a t show ?thesis using assms by simp + next + case (MINUS x91 x92) + with a t show ?thesis using assms by simp + next + case (EQUAL x101 x102) + with a t show ?thesis using assms by simp + next + case (LESS x111 x112) + with a t show ?thesis using assms by simp + next + case (AND x121 x122) + with a t show ?thesis using assms by simp + next + case (OR x131 x132) + with a t show ?thesis using assms by simp + next + case (NOT x131) + with a t show ?thesis using assms by simp + next + case (CALL x181 x182) + with a t show ?thesis using assms by simp + next + case (ECALL x191 x192 x193) + with a t show ?thesis using assms by simp + next + case CONTRACTS + with a t show ?thesis using assms by simp + qed + next + case f: FALSE + show ?thesis + proof (cases "eupdate e2") + case (INT x11 x12) + with a f show ?thesis using assms by simp + next + case (UINT x21 x22) + with a f show ?thesis using assms by simp + next + case (ADDRESS x3) + with a f show ?thesis using assms by simp + next + case (BALANCE x4) + with a f show ?thesis using assms by simp + next + case THIS + with a f show ?thesis using assms by simp + next + case SENDER + with a f show ?thesis using assms by simp + next + case VALUE + with a f show ?thesis using assms by simp + next + case TRUE + with a f show ?thesis using assms by simp + next + case FALSE + with a f show ?thesis using assms by simp + next + case (LVAL x7) + with a f show ?thesis using assms by simp + next + case (PLUS x81 x82) + with a f show ?thesis using assms by simp + next + case (MINUS x91 x92) + with a f show ?thesis using assms by simp + next + case (EQUAL x101 x102) + with a f show ?thesis using assms by simp + next + case (LESS x111 x112) + with a f show ?thesis using assms by simp + next + case (AND x121 x122) + with a f show ?thesis using assms by simp + next + case (OR x131 x132) + with a f show ?thesis using assms by simp + next + case (NOT x131) + with a f show ?thesis using assms by simp + next + case (CALL x181 x182) + with a f show ?thesis using assms by simp + next + case (ECALL x191 x192 x193) + with a f show ?thesis using assms by simp + next + case CONTRACTS + with a f show ?thesis using assms by simp + qed + next + case (LVAL x7) + with a show ?thesis using assms by simp + next + case (PLUS x81 x82) + with a show ?thesis using assms by simp + next + case (MINUS x91 x92) + with a show ?thesis using assms by simp + next + case (EQUAL x101 x102) + with a show ?thesis using assms by simp + next + case (LESS x111 x112) + with a show ?thesis using assms by simp + next + case (AND x121 x122) + with a show ?thesis using assms by simp + next + case (OR x131 x132) + with a show ?thesis using assms by simp + next + case (NOT x131) + with a show ?thesis using assms by simp + next + case (CALL x181 x182) + with a show ?thesis using assms by simp + next + case (ECALL x191 x192 x193) + with a show ?thesis using assms by simp + next + case CONTRACTS + with a show ?thesis using assms by simp + qed +next + case o: (OR e1 e2) + show ?thesis + proof (cases "eupdate e1") + case (INT x11 x12) + with o show ?thesis using assms by simp + next + case (UINT x21 x22) + with o show ?thesis using assms by simp + next + case (ADDRESS x3) + with o show ?thesis using assms by simp + next + case (BALANCE x4) + with o show ?thesis using assms by simp + next + case THIS + with o show ?thesis using assms by simp + next + case SENDER + with o show ?thesis using assms by simp + next + case VALUE + with o show ?thesis using assms by simp + next + case t: TRUE + show ?thesis + proof (cases "eupdate e2") + case (INT x11 x12) + with o t show ?thesis using assms by simp + next + case (UINT x21 x22) + with o t show ?thesis using assms by simp + next + case (ADDRESS x3) + with o t show ?thesis using assms by simp + next + case (BALANCE x4) + with o t show ?thesis using assms by simp + next + case THIS + with o t show ?thesis using assms by simp + next + case SENDER + with o t show ?thesis using assms by simp + next + case VALUE + with o t show ?thesis using assms by simp + next + case TRUE + with o t show ?thesis using assms by simp + next + case FALSE + with o t show ?thesis using assms by simp + next + case (LVAL x7) + with o t show ?thesis using assms by simp + next + case (PLUS x81 x82) + with o t show ?thesis using assms by simp + next + case (MINUS x91 x92) + with o t show ?thesis using assms by simp + next + case (EQUAL x101 x102) + with o t show ?thesis using assms by simp + next + case (LESS x111 x112) + with o t show ?thesis using assms by simp + next + case (AND x121 x122) + with o t show ?thesis using assms by simp + next + case (OR x131 x132) + with o t show ?thesis using assms by simp + next + case (NOT x131) + with o t show ?thesis using assms by simp + next + case (CALL x181 x182) + with o t show ?thesis using assms by simp + next + case (ECALL x191 x192 x193) + with o t show ?thesis using assms by simp + next + case CONTRACTS + with o t show ?thesis using assms by simp + qed + next + case f: FALSE + show ?thesis + proof (cases "eupdate e2") + case (INT x11 x12) + with o f show ?thesis using assms by simp + next + case (UINT x21 x22) + with o f show ?thesis using assms by simp + next + case (ADDRESS x3) + with o f show ?thesis using assms by simp + next + case (BALANCE x4) + with o f show ?thesis using assms by simp + next + case THIS + with o f show ?thesis using assms by simp + next + case SENDER + with o f show ?thesis using assms by simp + next + case VALUE + with o f show ?thesis using assms by simp + next + case TRUE + with o f show ?thesis using assms by simp + next + case FALSE + with o f show ?thesis using assms by simp + next + case (LVAL x7) + with o f show ?thesis using assms by simp + next + case (PLUS x81 x82) + with o f show ?thesis using assms by simp + next + case (MINUS x91 x92) + with o f show ?thesis using assms by simp + next + case (EQUAL x101 x102) + with o f show ?thesis using assms by simp + next + case (LESS x111 x112) + with o f show ?thesis using assms by simp + next + case (AND x121 x122) + with o f show ?thesis using assms by simp + next + case (OR x131 x132) + with o f show ?thesis using assms by simp + next + case (NOT x131) + with o f show ?thesis using assms by simp + next + case (CALL x181 x182) + with o f show ?thesis using assms by simp + next + case (ECALL x191 x192 x193) + with o f show ?thesis using assms by simp + next + case CONTRACTS + with o f show ?thesis using assms by simp + qed + next + case (LVAL x7) + with o show ?thesis using assms by simp + next + case (PLUS x81 x82) + with o show ?thesis using assms by simp + next + case (MINUS x91 x92) + with o show ?thesis using assms by simp + next + case (EQUAL x101 x102) + with o show ?thesis using assms by simp + next + case (LESS x111 x112) + with o show ?thesis using assms by simp + next + case (AND x121 x122) + with o show ?thesis using assms by simp + next + case (OR x131 x132) + with o show ?thesis using assms by simp + next + case (NOT x131) + with o show ?thesis using assms by simp + next + case (CALL x181 x182) + with o show ?thesis using assms by simp + next + case (ECALL x191 x192 x193) + with o show ?thesis using assms by simp + next + case CONTRACTS + with o show ?thesis using assms by simp + qed +next + case o: (NOT x) + show ?thesis + proof (cases "eupdate x") + case (INT x11 x12) + with o show ?thesis using assms by simp + next + case (UINT x21 x22) + with o show ?thesis using assms by simp + next + case (ADDRESS x3) + with o show ?thesis using assms by simp + next + case (BALANCE x4) + with o show ?thesis using assms by simp + next + case THIS + with o show ?thesis using assms by simp + next + case SENDER + with o show ?thesis using assms by simp + next + case VALUE + with o show ?thesis using assms by simp + next + case t: TRUE + with o show ?thesis using assms by simp + next + case f: FALSE + with o show ?thesis using assms by simp + next + case (LVAL x7) + with o show ?thesis using assms by simp + next + case (PLUS x81 x82) + with o show ?thesis using assms by simp + next + case (MINUS x91 x92) + with o show ?thesis using assms by simp + next + case (EQUAL x101 x102) + with o show ?thesis using assms by simp + next + case (LESS x111 x112) + with o show ?thesis using assms by simp + next + case (AND x121 x122) + with o show ?thesis using assms by simp + next + case (OR x131 x132) + with o show ?thesis using assms by simp + next + case (NOT x131) + with o show ?thesis using assms by simp + next + case (CALL x181 x182) + with o show ?thesis using assms by simp + next + case (ECALL x191 x192 x193) + with o show ?thesis using assms by simp + next + case CONTRACTS + with o show ?thesis using assms by simp + qed +next + case (CALL x181 x182) + with assms show ?thesis by simp +next + case (ECALL x191 x192 x193) + with assms show ?thesis by simp +next + case CONTRACTS + with assms show ?thesis by simp +qed + +lemma no_gas: + assumes "\ g > costs_ex ex env cd st" + shows "expr ex env cd st g = Exception Gas" +proof (cases ex) + case (INT x11 x12) + with assms show ?thesis by (simp add: Statements.solidity.expr.simps) +next + case (UINT x21 x22) + with assms show ?thesis by (simp add: Statements.solidity.expr.simps) +next + case (ADDRESS x3) + with assms show ?thesis by (simp add: Statements.solidity.expr.simps) +next + case (BALANCE x4) + with assms show ?thesis by (simp add: Statements.solidity.expr.simps) +next + case THIS + with assms show ?thesis by (simp add: Statements.solidity.expr.simps) +next + case SENDER + with assms show ?thesis by (simp add: Statements.solidity.expr.simps) +next + case VALUE + with assms show ?thesis by (simp add: Statements.solidity.expr.simps) +next + case TRUE + with assms show ?thesis by (simp add: Statements.solidity.expr.simps) +next + case FALSE + with assms show ?thesis by (simp add: Statements.solidity.expr.simps) +next + case (LVAL x10) + with assms show ?thesis by (simp add: Statements.solidity.expr.simps) +next + case (PLUS x111 x112) + with assms show ?thesis by (simp add: Statements.solidity.expr.simps) +next + case (MINUS x121 x122) + with assms show ?thesis by (simp add: Statements.solidity.expr.simps) +next + case (EQUAL x131 x132) + with assms show ?thesis by (simp add: Statements.solidity.expr.simps) +next + case (LESS x141 x142) + with assms show ?thesis by (simp add: Statements.solidity.expr.simps) +next + case (AND x151 x152) + with assms show ?thesis by (simp add: Statements.solidity.expr.simps) +next + case (OR x161 x162) + with assms show ?thesis by (simp add: Statements.solidity.expr.simps) +next + case (NOT x17) + with assms show ?thesis by (simp add: Statements.solidity.expr.simps) +next + case (CALL x181 x182) + with assms show ?thesis by (simp add: Statements.solidity.expr.simps) +next + case (ECALL x191 x192 x193) + with assms show ?thesis by (simp add: Statements.solidity.expr.simps) +next + case CONTRACTS + with assms show ?thesis by (simp add: Statements.solidity.expr.simps) +qed + +lemma lift_eq: + assumes "expr e1 env cd st g = expr e1' env cd st g" + and "\g' rv. expr e1 env cd st g = Normal (rv, g') \ expr e2 env cd st g'= expr e2' env cd st g'" + shows "lift expr f e1 e2 env cd st g =lift expr f e1' e2' env cd st g" +proof (cases "expr e1 env cd st g") + case s1: (n a g') + then show ?thesis + proof (cases a) + case f1:(Pair a b) + then show ?thesis + proof (cases a) + case k1:(KValue x1) + then show ?thesis + proof (cases b) + case v1: (Value x1) + then show ?thesis + proof (cases "expr e2 env cd st g'") + case s2: (n a' g'') + then show ?thesis + proof (cases a') + case f2:(Pair a' b') + then show ?thesis + proof (cases a') + case (KValue x1') + with s1 f1 k1 v1 assms(1) assms(2) show ?thesis by auto + next + case (KCDptr x2) + with s1 f1 k1 v1 assms(1) assms(2) show ?thesis by auto + next + case (KMemptr x2') + with s1 f1 k1 v1 assms(1) assms(2) show ?thesis by auto + next + case (KStoptr x3') + with s1 f1 k1 v1 assms(1) assms(2) show ?thesis by auto + qed + qed + next + case (e e) + then show ?thesis using k1 s1 v1 assms(1) assms(2) f1 by auto + qed + next + case (Calldata x2) + then show ?thesis using k1 s1 assms(1) f1 by auto + next + case (Memory x2) + then show ?thesis using k1 s1 assms(1) f1 by auto + next + case (Storage x3) + then show ?thesis using k1 s1 assms(1) f1 by auto + qed + next + case (KCDptr x2) + then show ?thesis using s1 assms(1) f1 by fastforce + next + case (KMemptr x2) + then show ?thesis using s1 assms(1) f1 by fastforce + next + case (KStoptr x3) + then show ?thesis using s1 assms(1) f1 by fastforce + qed + qed +next + case (e e) + then show ?thesis using assms(1) by simp +qed + +lemma ssel_eq_ssel: + "(\i g. i \ set ix \ expr i env cd st g = expr (f i) env cd st g) + \ ssel tp loc ix env cd st g = ssel tp loc (map f ix) env cd st g" +proof (induction ix arbitrary: tp loc env cd st g) + case Nil + then show ?case by simp +next + case c1: (Cons i ix) + then show ?case + proof (cases tp) + case tp1: (STArray al tp) + then show ?thesis + proof (cases "expr i env cd st g") + case s1: (n a g') + then show ?thesis + proof (cases a) + case f1: (Pair a b) + then show ?thesis + proof (cases a) + case k1: (KValue v) + then show ?thesis + proof (cases b) + case v1: (Value t) + then show ?thesis + proof (cases "less t (TUInt 256) v (ShowL\<^sub>i\<^sub>n\<^sub>t al)") + case None + with v1 k1 tp1 s1 c1.prems f1 show ?thesis by (simp add:Statements.solidity.ssel.simps) + next + case s2: (Some a) + then show ?thesis + proof (cases a) + case p1: (Pair a b) + then show ?thesis + proof (cases b) + case (TSInt x1) + with s2 p1 v1 k1 tp1 s1 c1.prems f1 show ?thesis by (simp add:Statements.solidity.ssel.simps) + next + case (TUInt x2) + with s2 p1 v1 k1 tp1 s1 c1.prems f1 show ?thesis by (simp add:Statements.solidity.ssel.simps) + next + case b1: TBool + show ?thesis + proof cases + assume "a = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True" + from c1.IH[OF c1.prems] have + "ssel tp (hash loc v) ix env cd st g' = ssel tp (hash loc v) (map f ix) env cd st g'" + by simp + with mp s2 b1 p1 v1 k1 tp1 s1 c1.prems f1 show ?thesis by (simp add:Statements.solidity.ssel.simps) + next + assume "\ a = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True" + with s2 p1 v1 k1 tp1 s1 c1.prems f1 show ?thesis by (simp add:Statements.solidity.ssel.simps) + qed + next + case TAddr + with s2 p1 v1 k1 tp1 s1 c1.prems f1 show ?thesis by (simp add:Statements.solidity.ssel.simps) + qed + qed + qed + next + case (Calldata x2) + with k1 tp1 s1 c1.prems f1 show ?thesis by (simp add:Statements.solidity.ssel.simps) + next + case (Memory x2) + with k1 tp1 s1 c1.prems f1 show ?thesis by (simp add:Statements.solidity.ssel.simps) + next + case (Storage x3) + with k1 tp1 s1 c1.prems f1 show ?thesis by (simp add:Statements.solidity.ssel.simps) + qed + next + case (KCDptr x2) + with tp1 s1 c1.prems f1 show ?thesis by (simp add:Statements.solidity.ssel.simps) + next + case (KMemptr x2) + with tp1 s1 c1.prems f1 show ?thesis by (simp add:Statements.solidity.ssel.simps) + next + case (KStoptr x3) + with tp1 s1 c1.prems f1 show ?thesis by (simp add:Statements.solidity.ssel.simps) + qed + qed + next + case (e e) + with tp1 c1.prems show ?thesis by (simp add:Statements.solidity.ssel.simps) + qed + next + case tp1: (STMap _ t) + then show ?thesis + proof (cases "expr i env cd st g") + case s1: (n a g') + then show ?thesis + proof (cases a) + case f1: (Pair a b) + then show ?thesis + proof (cases a) + case k1: (KValue v) + from c1.IH[OF c1.prems] have + "ssel tp (hash loc v) ix env cd st g = ssel tp (hash loc v) (map f ix) env cd st g" by simp + with k1 tp1 s1 c1 f1 show ?thesis by (simp add:Statements.solidity.ssel.simps) + next + case (KCDptr x2) + with tp1 s1 c1.prems f1 show ?thesis by (simp add:Statements.solidity.ssel.simps) + next + case (KMemptr x2) + with tp1 s1 c1.prems f1 show ?thesis by (simp add:Statements.solidity.ssel.simps) + next + case (KStoptr x3) + with tp1 s1 c1.prems f1 show ?thesis by (simp add:Statements.solidity.ssel.simps) + qed + qed + next + case (e e) + with tp1 c1.prems show ?thesis by (simp add:Statements.solidity.ssel.simps) + qed + next + case (STValue x2) + then show ?thesis by (simp add:Statements.solidity.ssel.simps) + qed +qed + +lemma msel_eq_msel: +"(\i g. i \ set ix \ expr i env cd st g = expr (f i) env cd st g) \ + msel c tp loc ix env cd st g = msel c tp loc (map f ix) env cd st g" +proof (induction ix arbitrary: c tp loc env cd st g) + case Nil + then show ?case by simp +next + case c1: (Cons i ix) + then show ?case + proof (cases tp) + case tp1: (MTArray al tp) + then show ?thesis + proof (cases ix) + case Nil + thus ?thesis using tp1 c1.prems by (auto simp add:Statements.solidity.msel.simps) + next + case c2: (Cons a list) + then show ?thesis + proof (cases "expr i env cd st g") + case s1: (n a g') + then show ?thesis + proof (cases a) + case f1: (Pair a b) + then show ?thesis + proof (cases a) + case k1: (KValue v) + then show ?thesis + proof (cases b) + case v1: (Value t) + then show ?thesis + proof (cases "less t (TUInt 256) v (ShowL\<^sub>i\<^sub>n\<^sub>t al)") + case None + with v1 k1 tp1 s1 c1.prems f1 show ?thesis using c2 by (simp add:Statements.solidity.msel.simps) + next + case s2: (Some a) + then show ?thesis + proof (cases a) + case p1: (Pair a b) + then show ?thesis + proof (cases b) + case (TSInt x1) + with s2 p1 v1 k1 tp1 s1 c1.prems f1 show ?thesis using c2 by (simp add:Statements.solidity.msel.simps) + next + case (TUInt x2) + with s2 p1 v1 k1 tp1 s1 c1.prems f1 show ?thesis using c2 by (simp add:Statements.solidity.msel.simps) + next + case b1: TBool + show ?thesis + proof cases + assume "a = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True" + then show ?thesis + proof (cases c) + case True + then show ?thesis + proof (cases "accessStore (hash loc v) (memory st)") + case None + with s2 b1 p1 v1 k1 tp1 s1 c1.prems f1 True show ?thesis using c2 by (simp add:Statements.solidity.msel.simps) + next + case s3: (Some a) + then show ?thesis + proof (cases a) + case (MValue x1) + with s2 s3 b1 p1 v1 k1 tp1 s1 c1.prems f1 True show ?thesis using c2 by (simp add:Statements.solidity.msel.simps) + next + case mp: (MPointer l) + from c1.IH[OF c1.prems] + have "msel c tp l ix env cd st g' = msel c tp l (map f ix) env cd st g'" by simp + with mp s2 s3 b1 p1 v1 k1 tp1 s1 c1.prems f1 True show ?thesis using c2 by (simp add:Statements.solidity.msel.simps) + qed + qed + next + case False + then show ?thesis + proof (cases "accessStore (hash loc v) cd") + case None + with s2 b1 p1 v1 k1 tp1 s1 c1.prems f1 False show ?thesis using c2 by (simp add:Statements.solidity.msel.simps) + next + case s3: (Some a) + then show ?thesis + proof (cases a) + case (MValue x1) + with s2 s3 b1 p1 v1 k1 tp1 s1 c1.prems f1 False show ?thesis using c2 by (simp add:Statements.solidity.msel.simps) + next + case mp: (MPointer l) + from c1.IH[OF c1.prems] + have "msel c tp l ix env cd st g' = msel c tp l (map f ix) env cd st g'" by simp + with mp s2 s3 b1 p1 v1 k1 tp1 s1 c1.prems f1 False show ?thesis using c2 by (simp add:Statements.solidity.msel.simps) + qed + qed + qed + next + assume "\ a = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True" + with s2 p1 v1 k1 tp1 s1 c1.prems f1 show ?thesis using c2 by (simp add:Statements.solidity.msel.simps) + qed + next + case TAddr + with s2 p1 v1 k1 tp1 s1 c1.prems f1 show ?thesis using c2 by (simp add:Statements.solidity.msel.simps) + qed + qed + qed + next + case (Calldata x2) + with k1 tp1 s1 c1.prems f1 show ?thesis using c2 by (simp add:Statements.solidity.msel.simps) + next + case (Memory x2) + with k1 tp1 s1 c1.prems f1 show ?thesis using c2 by (simp add:Statements.solidity.msel.simps) + next + case (Storage x3) + with k1 tp1 s1 c1.prems f1 show ?thesis using c2 by (simp add:Statements.solidity.msel.simps) + qed + next + case (KCDptr x2) + with tp1 s1 c1.prems f1 show ?thesis using c2 by (simp add:Statements.solidity.msel.simps) + next + case (KMemptr x2) + with tp1 s1 c1.prems f1 show ?thesis using c2 by (simp add:Statements.solidity.msel.simps) + next + case (KStoptr x3) + with tp1 s1 c1.prems f1 show ?thesis using c2 by (simp add:Statements.solidity.msel.simps) + qed + qed + next + case (e e) + with tp1 c1.prems show ?thesis using c2 by (simp add:Statements.solidity.msel.simps) + qed + qed + next + case (MTValue x2) + then show ?thesis by (simp add:Statements.solidity.msel.simps) + qed +qed + +lemma ref_eq: + assumes "\e g. e \ set ex \ expr e env cd st g = expr (f e) env cd st g" + shows "rexp (Ref i ex) env cd st g = rexp (Ref i (map f ex)) env cd st g" +proof (cases "fmlookup (denvalue env) i") + case None + then show ?thesis by (simp add:Statements.solidity.rexp.simps) +next + case s1: (Some a) + then show ?thesis + proof (cases a) + case p1: (Pair tp b) + then show ?thesis + proof (cases b) + case k1: (Stackloc l) + then show ?thesis + proof (cases "accessStore l (stack st)") + case None + with s1 p1 k1 show ?thesis by (simp add:Statements.solidity.rexp.simps) + next + case s2: (Some a') + then show ?thesis + proof (cases a') + case (KValue _) + with s1 s2 p1 k1 show ?thesis by (simp add:Statements.solidity.rexp.simps) + next + case cp: (KCDptr cp) + then show ?thesis + proof (cases tp) + case (Value x1) + with cp s1 s2 p1 k1 show ?thesis by (simp add:Statements.solidity.rexp.simps) + next + case mt: (Calldata ct) + from msel_eq_msel have + "msel False ct cp ex env cd st=msel False ct cp (map f ex) env cd st" using assms by blast + thus ?thesis using s1 s2 p1 k1 mt cp by (simp add:Statements.solidity.rexp.simps) + next + case mt: (Memory mt) + from msel_eq_msel have + "msel True mt cp ex env cd st=msel True mt cp (map f ex) env cd st" using assms by blast + thus ?thesis using s1 s2 p1 k1 mt cp by (simp add:Statements.solidity.rexp.simps) + next + case (Storage x3) + with cp s1 s2 p1 k1 show ?thesis by (simp add:Statements.solidity.rexp.simps) + qed + next + case mp: (KMemptr mp) + then show ?thesis + proof (cases tp) + case (Value x1) + with mp s1 s2 p1 k1 show ?thesis by (simp add:Statements.solidity.rexp.simps) + next + case mt: (Calldata ct) + from msel_eq_msel have + "msel True ct mp ex env cd st=msel True ct mp (map f ex) env cd st" using assms by blast + thus ?thesis using s1 s2 p1 k1 mt mp by (simp add:Statements.solidity.rexp.simps) + next + case mt: (Memory mt) + from msel_eq_msel have + "msel True mt mp ex env cd st=msel True mt mp (map f ex) env cd st" using assms by blast + thus ?thesis using s1 s2 p1 k1 mt mp by (simp add:Statements.solidity.rexp.simps) + next + case (Storage x3) + with mp s1 s2 p1 k1 show ?thesis by (simp add:Statements.solidity.rexp.simps) + qed + next + case sp: (KStoptr sp) + then show ?thesis + proof (cases tp) + case (Value x1) + then show ?thesis using s1 s2 p1 k1 sp by (simp add:Statements.solidity.rexp.simps) + next + case (Calldata x2) + then show ?thesis using s1 s2 p1 k1 sp by (simp add:Statements.solidity.rexp.simps) + next + case (Memory x2) + then show ?thesis using s1 s2 p1 k1 sp by (simp add:Statements.solidity.rexp.simps) + next + case st: (Storage stp) + from ssel_eq_ssel have + "ssel stp sp ex env cd st=ssel stp sp (map f ex) env cd st" using assms by blast + thus ?thesis using s1 s2 p1 k1 st sp by (simp add:Statements.solidity.rexp.simps) + qed + qed + qed + next + case sl:(Storeloc sl) + then show ?thesis + proof (cases tp) + case (Value x1) + then show ?thesis using s1 p1 sl by (simp add:Statements.solidity.rexp.simps) + next + case (Calldata x2) + then show ?thesis using s1 p1 sl by (simp add:Statements.solidity.rexp.simps) + next + case (Memory x2) + then show ?thesis using s1 p1 sl by (simp add:Statements.solidity.rexp.simps) + next + case st: (Storage stp) + from ssel_eq_ssel have + "ssel stp sl ex env cd st=ssel stp sl (map f ex) env cd st" using assms by blast + thus ?thesis using s1 sl p1 st by (simp add:Statements.solidity.rexp.simps) + qed + qed + qed +qed + +text\ + The following theorem proves that the update function preserves the semantics of expressions. +\ +theorem update_correctness: + "\g. expr ex env cd st g = expr (eupdate ex) env cd st g" + "\g. rexp lv env cd st g = rexp (lupdate lv) env cd st g" +proof (induction ex and lv) + case (Id x g) + then show ?case by simp +next + case (Ref d ix g) + then show ?case using ref_eq[where f="eupdate"] by simp +next + case (INT b v g) + then show ?case + proof (cases "g > 0") + case True + then show ?thesis + proof cases + assume "b\vbits" + show ?thesis + proof cases + let ?m_def = "(-(2^(b-1)) + (v+2^(b-1)) mod (2^b))" + assume "v \ 0" + + from `b\vbits` True have + "expr (E.INT b v) env cd st g = Normal ((KValue (createSInt b v), Value (TSInt b)), g)" + by (simp add:Statements.solidity.expr.simps) + also have "createSInt b v = createSInt b ?m_def" using `b\vbits` `v \ 0` unfolding createSInt_def by auto + also from `v \ 0` `b\vbits` True have + "Normal ((KValue (createSInt b ?m_def), Value (TSInt b)),g) = expr (eupdate (E.INT b v)) env cd st g" + by (simp add:Statements.solidity.expr.simps) + finally show "expr (E.INT b v) env cd st g = expr (eupdate (E.INT b v)) env cd st g" by simp + next + let ?m_def = "(2^(b-1) - (-v+2^(b-1)-1) mod (2^b) - 1)" + assume "\ v \ 0" + + from `b\vbits` True have + "expr (E.INT b v) env cd st g = Normal ((KValue (createSInt b v), Value (TSInt b)), g)" by (simp add:Statements.solidity.expr.simps) + also have "createSInt b v = createSInt b ?m_def" using `b\vbits` `\ v \ 0` unfolding createSInt_def by auto + also from `\ v \ 0` `b\vbits` True have + "Normal ((KValue (createSInt b ?m_def), Value (TSInt b)),g) =expr (eupdate (E.INT b v)) env cd st g" + by (simp add:Statements.solidity.expr.simps) + finally show "expr (E.INT b v) env cd st g = expr (eupdate (E.INT b v)) env cd st g" by simp + qed + next + assume "\ b\vbits" + thus ?thesis by auto + qed + next + case False + then show ?thesis using no_gas by simp + qed +next + case (UINT x1 x2) + then show ?case by (simp add:Statements.solidity.expr.simps createUInt_def) +next + case (ADDRESS x) + then show ?case by simp +next + case (BALANCE x) + then show ?case by simp +next + case THIS + then show ?case by simp +next + case SENDER + then show ?case by simp +next + case VALUE + then show ?case by simp +next + case TRUE + then show ?case by simp +next + case FALSE + then show ?case by simp +next + case (LVAL x) + then show ?case by (simp add:Statements.solidity.expr.simps createUInt_def) +next + case p: (PLUS e1 e2 g) + show ?case + proof (cases "eupdate e1") + case i: (INT b1 v1) + with p.IH have expr1: "expr e1 env cd st g = expr (E.INT b1 v1) env cd st g" by simp + then show ?thesis + proof (cases "g > 0") + case True + then show ?thesis + proof (cases) + assume "b1 \ vbits" + with expr1 True + have "expr e1 env cd st g = Normal ((KValue (createSInt b1 v1), Value (TSInt b1)), g)" by (simp add:Statements.solidity.expr.simps createSInt_def) + moreover from i `b1 \ vbits` + have "v1 < 2^(b1-1)" and "v1 \ -(2^(b1-1))" using update_bounds_int by auto + moreover from `b1 \ vbits` have "0 < b1" by auto + ultimately have r1: "expr e1 env cd st g = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t v1), Value (TSInt b1)),g)" + using createSInt_id[of v1 b1] by simp + thus ?thesis + proof (cases "eupdate e2") + case i2: (INT b2 v2) + with p.IH have expr2: "expr e2 env cd st g = expr (E.INT b2 v2) env cd st g" by simp + then show ?thesis + proof (cases) + let ?v="v1 + v2" + assume "b2 \ vbits" + with expr2 True + have "expr e2 env cd st g = Normal ((KValue (createSInt b2 v2), Value (TSInt b2)),g)" by (simp add:Statements.solidity.expr.simps) + moreover from i2 `b2 \ vbits` + have "v2 < 2^(b2-1)" and "v2 \ -(2^(b2-1))" using update_bounds_int by auto + moreover from `b2 \ vbits` have "0 < b2" by auto + ultimately have r2: "expr e2 env cd st g = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t v2), Value (TSInt b2)),g)" + using createSInt_id[of v2 b2] by simp + thus ?thesis + proof (cases) + let ?x="- (2 ^ (max b1 b2 - 1)) + (?v + 2 ^ (max b1 b2 - 1)) mod 2 ^ max b1 b2" + assume "?v\0" + hence "createSInt (max b1 b2) ?v = (ShowL\<^sub>i\<^sub>n\<^sub>t ?x)" by (simp add: createSInt_def) + moreover have "add (TSInt b1) (TSInt b2) (ShowL\<^sub>i\<^sub>n\<^sub>t v1) (ShowL\<^sub>i\<^sub>n\<^sub>t v2) + = Some (createSInt (max b1 b2) ?v, TSInt (max b1 b2))" + using Read_ShowL_id add_def olift.simps(1)[of "(+)" b1 b2] by simp + ultimately have "expr (PLUS e1 e2) env cd st g + = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TSInt (max b1 b2))),g)" using r1 r2 True by (simp add:Statements.solidity.expr.simps) + moreover have "expr (eupdate (PLUS e1 e2)) env cd st g + = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TSInt (max b1 b2))),g)" + proof - + from `b1 \ vbits` `b2 \ vbits` `?v\0` + have "eupdate (PLUS e1 e2) = E.INT (max b1 b2) ?x" using i i2 by simp + moreover have "expr (E.INT (max b1 b2) ?x) env cd st g + = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TSInt (max b1 b2))),g)" + proof - + from `b1 \ vbits` `b2 \ vbits` have "max b1 b2 \ vbits" using vbits_max by simp + with True have "expr (E.INT (max b1 b2) ?x) env cd st g + = Normal ((KValue (createSInt (max b1 b2) ?x), Value (TSInt (max b1 b2))),g)" by (simp add:Statements.solidity.expr.simps) + moreover from `0 < b1` + have "?x < 2 ^ (max b1 b2 - 1)" using upper_bound3 by simp + moreover from `0 < b1` have "0 < max b1 b2" using max_def by simp + ultimately show ?thesis using createSInt_id[of ?x "max b1 b2"] by simp + qed + ultimately show ?thesis by simp + qed + ultimately show ?thesis by simp + next + let ?x="2^(max b1 b2 -1) - (-?v+2^(max b1 b2-1)-1) mod (2^max b1 b2) - 1" + assume "\ ?v\0" + hence "createSInt (max b1 b2) ?v = (ShowL\<^sub>i\<^sub>n\<^sub>t ?x)" unfolding createSInt_def by simp + moreover have "add (TSInt b1) (TSInt b2) (ShowL\<^sub>i\<^sub>n\<^sub>t v1) (ShowL\<^sub>i\<^sub>n\<^sub>t v2) + = Some (createSInt (max b1 b2) ?v, TSInt (max b1 b2))" + using Read_ShowL_id add_def olift.simps(1)[of "(+)" b1 b2] by simp + ultimately have "expr (PLUS e1 e2) env cd st g + = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TSInt (max b1 b2))),g)" using True r1 r2 by (simp add:Statements.solidity.expr.simps) + moreover have "expr (eupdate (PLUS e1 e2)) env cd st g + = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TSInt (max b1 b2))),g)" + proof - + from `b1 \ vbits` `b2 \ vbits` `\?v\0` + have "eupdate (PLUS e1 e2) = E.INT (max b1 b2) ?x" using i i2 by simp + moreover have "expr (E.INT (max b1 b2) ?x) env cd st g + = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TSInt (max b1 b2))),g)" + proof - + from `b1 \ vbits` `b2 \ vbits` have "max b1 b2 \ vbits" using vbits_max by simp + with True have "expr (E.INT (max b1 b2) ?x) env cd st g + = Normal ((KValue (createSInt (max b1 b2) ?x), Value (TSInt (max b1 b2))),g)" by (simp add:Statements.solidity.expr.simps) + moreover from `0 < b1` + have "?x \ - (2 ^ (max b1 b2 - 1))" using lower_bound2[of "max b1 b2" ?v] by simp + moreover from `b1 > 0` have "2^(max b1 b2 -1) > (0::nat)" by simp + hence "2^(max b1 b2 -1) - (-?v+2^(max b1 b2-1)-1) mod (2^max b1 b2) - 1 < 2 ^ (max b1 b2 - 1)" + by (simp add: algebra_simps flip: zle_diff1_eq) + moreover from `0 < b1` have "0 < max b1 b2" using max_def by simp + ultimately show ?thesis using createSInt_id[of ?x "max b1 b2"] by simp + qed + ultimately show ?thesis by simp + qed + ultimately show ?thesis by simp + qed + next + assume "\ b2 \ vbits" + with p i i2 show ?thesis by (simp add:Statements.solidity.expr.simps) + qed + next + case u2: (UINT b2 v2) + with p.IH have expr2: "expr e2 env cd st g = expr (UINT b2 v2) env cd st g" by simp + then show ?thesis + proof (cases) + let ?v="v1 + v2" + assume "b2 \ vbits" + with expr2 True + have "expr e2 env cd st g = Normal ((KValue (createUInt b2 v2), Value (TUInt b2)),g)" by (simp add:Statements.solidity.expr.simps) + moreover from u2 `b2 \ vbits` + have "v2 < 2^b2" and "v2 \ 0" using update_bounds_uint by auto + moreover from `b2 \ vbits` have "0 < b2" by auto + ultimately have r2: "expr e2 env cd st g = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t v2), Value (TUInt b2)),g)" + using createUInt_id[of v2 b2] by simp + thus ?thesis + proof (cases) + assume "b20" + hence "createSInt b1 ?v = (ShowL\<^sub>i\<^sub>n\<^sub>t ?x)" using `b2i\<^sub>n\<^sub>t v1) (ShowL\<^sub>i\<^sub>n\<^sub>t v2) + = Some (createSInt b1 ?v, TSInt b1)" + using Read_ShowL_id add_def olift.simps(3)[of "(+)" b1 b2] `b2i\<^sub>n\<^sub>t ?x), Value (TSInt b1)),g)" using r1 r2 True by (simp add:Statements.solidity.expr.simps) + moreover have "expr (eupdate (PLUS e1 e2)) env cd st g + = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TSInt b1)),g)" + proof - + from `b1 \ vbits` `b2 \ vbits` `?v\0` `b2i\<^sub>n\<^sub>t ?x), Value (TSInt b1)),g)" + proof - + from `b1 \ vbits` True have "expr (E.INT b1 ?x) env cd st g + = Normal ((KValue (createSInt b1 ?x), Value (TSInt b1)),g)" by (simp add:Statements.solidity.expr.simps) + moreover from `0 < b1` have "?x < 2 ^ (b1 - 1)" using upper_bound2 by simp + ultimately show ?thesis using createSInt_id[of ?x "b1"] `0 < b1` by simp + qed + ultimately show ?thesis by simp + qed + ultimately show ?thesis by simp + next + let ?x="2^(b1 -1) - (-?v+2^(b1-1)-1) mod (2^b1) - 1" + assume "\ ?v\0" + hence "createSInt b1 ?v = (ShowL\<^sub>i\<^sub>n\<^sub>t ?x)" unfolding createSInt_def by simp + moreover have "add (TSInt b1) (TUInt b2) (ShowL\<^sub>i\<^sub>n\<^sub>t v1) (ShowL\<^sub>i\<^sub>n\<^sub>t v2) + = Some (createSInt b1 ?v, TSInt b1)" + using Read_ShowL_id add_def olift.simps(3)[of "(+)" b1 b2] `b2i\<^sub>n\<^sub>t ?x), Value (TSInt b1)),g)" using r1 r2 True by (simp add:Statements.solidity.expr.simps) + moreover have "expr (eupdate (PLUS e1 e2)) env cd st g + = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TSInt b1)),g)" + proof - + from `b1 \ vbits` `b2 \ vbits` `\?v\0` `b2i\<^sub>n\<^sub>t ?x), Value (TSInt b1)),g)" + proof - + from `b1 \ vbits` True have "expr (E.INT b1 ?x) env cd st g + = Normal ((KValue (createSInt b1 ?x), Value (TSInt b1)),g)" by (simp add:Statements.solidity.expr.simps) + moreover from `0 < b1` have "?x \ - (2 ^ (b1 - 1))" using upper_bound2 by simp + moreover have "2^(b1-1) - (-?v+2^(b1-1)-1) mod (2^b1) - 1 < 2 ^ (b1 - 1)" + by (simp add: algebra_simps flip: int_one_le_iff_zero_less) + ultimately show ?thesis using createSInt_id[of ?x b1] `0 < b1` by simp + qed + ultimately show ?thesis by simp + qed + ultimately show ?thesis by simp + qed + next + assume "\ b2 < b1" + with i u2 have "eupdate (PLUS e1 e2) = PLUS (eupdate e1) (eupdate e2)" by simp + with p i show ?thesis by (simp add:Statements.solidity.expr.simps) + qed + next + assume "\ b2 \ vbits" + with p i u2 show ?thesis by (simp add:Statements.solidity.expr.simps) + qed + next + case (ADDRESS _) + with p i show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (BALANCE _) + with p i show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case THIS + with p i show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case SENDER + with p i show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case VALUE + with p i show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case TRUE + with p i show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case FALSE + with p i show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (LVAL _) + with p i show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (PLUS _ _) + with p i show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (MINUS _ _) + with p i show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (EQUAL _ _) + with p i show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (LESS _ _) + with p i show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (AND _ _) + with p i show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (OR _ _) + with p i show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (NOT _) + with p i show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (CALL x181 x182) + with p i show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (ECALL x191 x192 x193) + with p i show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case CONTRACTS + with p i show ?thesis by (simp add:Statements.solidity.expr.simps) + qed + next + assume "\ b1 \ vbits" + with p i show ?thesis by (simp add:Statements.solidity.expr.simps) + qed + next + case False + then show ?thesis using no_gas by simp + qed + next + case u: (UINT b1 v1) + with p.IH have expr1: "expr e1 env cd st g = expr (UINT b1 v1) env cd st g" by simp + then show ?thesis + proof (cases "g > 0") + case True + then show ?thesis + proof (cases) + assume "b1 \ vbits" + with expr1 True + have "expr e1 env cd st g = Normal ((KValue (createUInt b1 v1), Value (TUInt b1)),g)" by (simp add:Statements.solidity.expr.simps) + moreover from u `b1 \ vbits` + have "v1 < 2^b1" and "v1 \ 0" using update_bounds_uint by auto + moreover from `b1 \ vbits` have "0 < b1" by auto + ultimately have r1: "expr e1 env cd st g = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t v1), Value (TUInt b1)),g)" + by (simp add:Statements.solidity.expr.simps createUInt_def) + thus ?thesis + proof (cases "eupdate e2") + case u2: (UINT b2 v2) + with p.IH have expr2: "expr e2 env cd st g = expr (UINT b2 v2) env cd st g" by simp + then show ?thesis + proof (cases) + let ?v="v1 + v2" + let ?x="?v mod 2 ^ max b1 b2" + assume "b2 \ vbits" + with expr2 True + have "expr e2 env cd st g = Normal ((KValue (createUInt b2 v2), Value (TUInt b2)),g)" by (simp add:Statements.solidity.expr.simps) + moreover from u2 `b2 \ vbits` + have "v2 < 2^b2" and "v2 \ 0" using update_bounds_uint by auto + moreover from `b2 \ vbits` have "0 < b2" by auto + ultimately have r2: "expr e2 env cd st g = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t v2), Value (TUInt b2)),g)" + by (simp add:Statements.solidity.expr.simps createUInt_def) + moreover have "add (TUInt b1) (TUInt b2) (ShowL\<^sub>i\<^sub>n\<^sub>t v1) (ShowL\<^sub>i\<^sub>n\<^sub>t v2) + = Some (createUInt (max b1 b2) ?v, TUInt (max b1 b2))" + using Read_ShowL_id add_def olift.simps(2)[of "(+)" b1 b2] by simp + ultimately have "expr (PLUS e1 e2) env cd st g + = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TUInt (max b1 b2))),g)" using r1 True by (simp add:Statements.solidity.expr.simps createUInt_def) + moreover have "expr (eupdate (PLUS e1 e2)) env cd st g + = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TUInt (max b1 b2))),g)" + proof - + from `b1 \ vbits` `b2 \ vbits` + have "eupdate (PLUS e1 e2) = UINT (max b1 b2) ?x" using u u2 by simp + moreover have "expr (UINT (max b1 b2) ?x) env cd st g + = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TUInt (max b1 b2))),g)" + proof - + from `b1 \ vbits` `b2 \ vbits` have "max b1 b2 \ vbits" using vbits_max by simp + with True have "expr (UINT (max b1 b2) ?x) env cd st g + = Normal ((KValue (createUInt (max b1 b2) ?x), Value (TUInt (max b1 b2))),g)" by (simp add:Statements.solidity.expr.simps) + moreover from `0 < b1` + have "?x < 2 ^ (max b1 b2)" by simp + moreover from `0 < b1` have "0 < max b1 b2" using max_def by simp + ultimately show ?thesis by (simp add:Statements.solidity.expr.simps createUInt_def) + qed + ultimately show ?thesis by simp + qed + ultimately show ?thesis by simp + next + assume "\ b2 \ vbits" + with p u u2 show ?thesis by (simp add:Statements.solidity.expr.simps) + qed + next + case i2: (INT b2 v2) + with p.IH have expr2: "expr e2 env cd st g = expr (E.INT b2 v2) env cd st g" by simp + then show ?thesis + proof (cases) + let ?v="v1 + v2" + assume "b2 \ vbits" + with expr2 True + have "expr e2 env cd st g = Normal ((KValue (createSInt b2 v2), Value (TSInt b2)),g)" by (simp add:Statements.solidity.expr.simps) + moreover from i2 `b2 \ vbits` + have "v2 < 2^(b2-1)" and "v2 \ -(2^(b2-1))" using update_bounds_int by auto + moreover from `b2 \ vbits` have "0 < b2" by auto + ultimately have r2: "expr e2 env cd st g = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t v2), Value (TSInt b2)),g)" + using createSInt_id[of v2 b2] by simp + thus ?thesis + proof (cases) + assume "b10" + hence "createSInt b2 ?v = (ShowL\<^sub>i\<^sub>n\<^sub>t ?x)" using `b1i\<^sub>n\<^sub>t v1) (ShowL\<^sub>i\<^sub>n\<^sub>t v2) + = Some (createSInt b2 ?v, TSInt b2)" + using Read_ShowL_id add_def olift.simps(4)[of "(+)" b1 b2] `b1i\<^sub>n\<^sub>t ?x), Value (TSInt b2)),g)" using r1 r2 True by (simp add:Statements.solidity.expr.simps) + moreover have "expr (eupdate (PLUS e1 e2)) env cd st g + = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TSInt b2)),g)" + proof - + from `b1 \ vbits` `b2 \ vbits` `?v\0` `b1i\<^sub>n\<^sub>t ?x), Value (TSInt b2)),g)" + proof - + from `b2 \ vbits` True have "expr (E.INT b2 ?x) env cd st g + = Normal ((KValue (createSInt b2 ?x), Value (TSInt b2)),g)" by (simp add:Statements.solidity.expr.simps) + moreover from `0 < b2` have "?x < 2 ^ (b2 - 1)" using upper_bound2 by simp + ultimately show ?thesis using createSInt_id[of ?x "b2"] `0 < b2` by simp + qed + ultimately show ?thesis by simp + qed + ultimately show ?thesis by simp + next + let ?x="2^(b2 -1) - (-?v+2^(b2-1)-1) mod (2^b2) - 1" + assume "\ ?v\0" + hence "createSInt b2 ?v = (ShowL\<^sub>i\<^sub>n\<^sub>t ?x)" unfolding createSInt_def by simp + moreover have "add (TUInt b1) (TSInt b2) (ShowL\<^sub>i\<^sub>n\<^sub>t v1) (ShowL\<^sub>i\<^sub>n\<^sub>t v2) + = Some (createSInt b2 ?v, TSInt b2)" + using Read_ShowL_id add_def olift.simps(4)[of "(+)" b1 b2] `b1i\<^sub>n\<^sub>t ?x), Value (TSInt b2)),g)" using r1 r2 True by (simp add:Statements.solidity.expr.simps) + moreover have "expr (eupdate (PLUS e1 e2)) env cd st g + = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TSInt b2)),g)" + proof - + from `b1 \ vbits` `b2 \ vbits` `\?v\0` `b1i\<^sub>n\<^sub>t ?x), Value (TSInt b2)),g)" + proof - + from `b2 \ vbits` True have "expr (E.INT b2 ?x) env cd st g + = Normal ((KValue (createSInt b2 ?x), Value (TSInt b2)),g)" by (simp add:Statements.solidity.expr.simps) + moreover from `0 < b2` have "?x \ - (2 ^ (b2 - 1))" using upper_bound2 by simp + moreover have "2^(b2-1) - (-?v+2^(b2-1)-1) mod (2^b2) - 1 < 2 ^ (b2 - 1)" + by (simp add: algebra_simps flip: int_one_le_iff_zero_less) + ultimately show ?thesis using createSInt_id[of ?x b2] `0 < b2` by simp + qed + ultimately show ?thesis by simp + qed + ultimately show ?thesis by simp + qed + next + assume "\ b1 < b2" + with p u i2 show ?thesis by (simp add:Statements.solidity.expr.simps) + qed + next + assume "\ b2 \ vbits" + with p u i2 show ?thesis by (simp add:Statements.solidity.expr.simps) + qed + next + case (ADDRESS _) + with p u show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (BALANCE _) + with p u show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case THIS + with p u show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case SENDER + with p u show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case VALUE + with p u show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case TRUE + with p u show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case FALSE + with p u show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (LVAL _) + with p u show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (PLUS _ _) + with p u show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (MINUS _ _) + with p u show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (EQUAL _ _) + with p u show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (LESS _ _) + with p u show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (AND _ _) + with p u show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (OR _ _) + with p u show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (NOT _) + with p u show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (CALL x181 x182) + with p u show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (ECALL x191 x192 x193) + with p u show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case CONTRACTS + with p u show ?thesis by (simp add:Statements.solidity.expr.simps) + qed + next + assume "\ b1 \ vbits" + with p u show ?thesis by (simp add:Statements.solidity.expr.simps) + qed + next + case False + then show ?thesis using no_gas by simp + qed + next + case (ADDRESS x3) + with p show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (BALANCE x4) + with p show ?thesis using lift_eq[of e1 env cd st g "eupdate e1" e2 "eupdate e2"] by (auto simp add:Statements.solidity.expr.simps) + next + case THIS + with p show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case SENDER + with p show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case VALUE + with p show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case TRUE + with p show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case FALSE + with p show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (LVAL x7) + with p show ?thesis using lift_eq[of e1 env cd st g "eupdate e1" e2 "eupdate e2"] by (auto simp add:Statements.solidity.expr.simps) + next + case (PLUS x81 x82) + with p show ?thesis using lift_eq[of e1 env cd st g "eupdate e1" e2 "eupdate e2"] by (auto simp add:Statements.solidity.expr.simps) + next + case (MINUS x91 x92) + with p show ?thesis using lift_eq[of e1 env cd st g "eupdate e1" e2 "eupdate e2"] by (auto simp add:Statements.solidity.expr.simps) + next + case (EQUAL x101 x102) + with p show ?thesis using lift_eq[of e1 env cd st g "eupdate e1" e2 "eupdate e2"] by (auto simp add:Statements.solidity.expr.simps) + next + case (LESS x111 x112) + with p show ?thesis using lift_eq[of e1 env cd st g "eupdate e1" e2 "eupdate e2"] by (auto simp add:Statements.solidity.expr.simps) + next + case (AND x121 x122) + with p show ?thesis using lift_eq[of e1 env cd st g "eupdate e1" e2 "eupdate e2"] by (auto simp add:Statements.solidity.expr.simps) + next + case (OR x131 x132) + with p show ?thesis using lift_eq[of e1 env cd st g "eupdate e1" e2 "eupdate e2"] by (auto simp add:Statements.solidity.expr.simps) + next + case (NOT x131) + with p show ?thesis using lift_eq[of e1 env cd st g "eupdate e1" e2 "eupdate e2"] by (auto simp add:Statements.solidity.expr.simps) + next + case (CALL x181 x182) + with p show ?thesis using lift_eq[of e1 env cd st g "eupdate e1" e2 "eupdate e2"] by (auto simp add:Statements.solidity.expr.simps) + next + case (ECALL x191 x192 x193) + with p show ?thesis using lift_eq[of e1 env cd st g "eupdate e1" e2 "eupdate e2"] by (auto simp add:Statements.solidity.expr.simps) + next + case CONTRACTS + with p show ?thesis using lift_eq[of e1 env cd st g "eupdate e1" e2 "eupdate e2"] by (auto simp add:Statements.solidity.expr.simps) + qed +next + case m: (MINUS e1 e2 g) + show ?case + proof (cases "eupdate e1") + case i: (INT b1 v1) + with m.IH have expr1: "expr e1 env cd st g = expr (E.INT b1 v1) env cd st g" by simp + then show ?thesis + proof (cases "g > 0") + case True + show ?thesis + proof (cases) + assume "b1 \ vbits" + with expr1 True + have "expr e1 env cd st g = Normal ((KValue (createSInt b1 v1), Value (TSInt b1)),g)" by (simp add:Statements.solidity.expr.simps) + moreover from i `b1 \ vbits` + have "v1 < 2^(b1-1)" and "v1 \ -(2^(b1-1))" using update_bounds_int by auto + moreover from `b1 \ vbits` have "0 < b1" by auto + ultimately have r1: "expr e1 env cd st g = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t v1), Value (TSInt b1)),g)" + using createSInt_id[of v1 b1] by simp + thus ?thesis + proof (cases "eupdate e2") + case i2: (INT b2 v2) + with m.IH have expr2: "expr e2 env cd st g = expr (E.INT b2 v2) env cd st g" by simp + then show ?thesis + proof (cases) + let ?v="v1 - v2" + assume "b2 \ vbits" + with expr2 True + have "expr e2 env cd st g = Normal ((KValue (createSInt b2 v2), Value (TSInt b2)),g)" by (simp add:Statements.solidity.expr.simps) + moreover from i2 `b2 \ vbits` + have "v2 < 2^(b2-1)" and "v2 \ -(2^(b2-1))" using update_bounds_int by auto + moreover from `b2 \ vbits` have "0 < b2" by auto + ultimately have r2: "expr e2 env cd st g = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t v2), Value (TSInt b2)),g)" + using createSInt_id[of v2 b2] by simp + + from `b1 \ vbits` `b2 \ vbits` have + u_def: "eupdate (MINUS e1 e2) = + (let v = v1 - v2 + in if 0 \ v + then E.INT (max b1 b2) + (- (2 ^ (max b1 b2 - 1)) + (v + 2 ^ (max b1 b2 - 1)) mod 2 ^ max b1 b2) + else E.INT (max b1 b2) + (2 ^ (max b1 b2 - 1) - (- v + 2 ^ (max b1 b2 - 1) - 1) mod 2 ^ max b1 b2 - 1))" + using i i2 eupdate.simps(11)[of e1 e2] by simp + + show ?thesis + proof (cases) + let ?x="- (2 ^ (max b1 b2 - 1)) + (?v + 2 ^ (max b1 b2 - 1)) mod 2 ^ max b1 b2" + assume "?v\0" + hence "createSInt (max b1 b2) ?v = (ShowL\<^sub>i\<^sub>n\<^sub>t ?x)" unfolding createSInt_def by simp + moreover have "sub (TSInt b1) (TSInt b2) (ShowL\<^sub>i\<^sub>n\<^sub>t v1) (ShowL\<^sub>i\<^sub>n\<^sub>t v2) + = Some (createSInt (max b1 b2) ?v, TSInt (max b1 b2))" + using Read_ShowL_id sub_def olift.simps(1)[of "(-)" b1 b2] by simp + ultimately have "expr (MINUS e1 e2) env cd st g + = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TSInt (max b1 b2))),g)" using r1 r2 True by (simp add:Statements.solidity.expr.simps) + moreover have "expr (eupdate (MINUS e1 e2)) env cd st g + = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TSInt (max b1 b2))),g)" + proof - + from u_def have "eupdate (MINUS e1 e2) = E.INT (max b1 b2) ?x" using `?v\0` by simp + moreover have "expr (E.INT (max b1 b2) ?x) env cd st g + = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TSInt (max b1 b2))),g)" + proof - + from `b1 \ vbits` `b2 \ vbits` have "max b1 b2 \ vbits" using vbits_max by simp + with True have "expr (E.INT (max b1 b2) ?x) env cd st g + = Normal ((KValue (createSInt (max b1 b2) ?x), Value (TSInt (max b1 b2))),g)" by (simp add:Statements.solidity.expr.simps) + moreover from `0 < b1` + have "?x < 2 ^ (max b1 b2 - 1)" using upper_bound2 by simp + moreover from `0 < b1` have "0 < max b1 b2" using max_def by simp + ultimately show ?thesis using createSInt_id[of ?x "max b1 b2"] by simp + qed + ultimately show ?thesis by simp + qed + ultimately show ?thesis by simp + next + let ?x="2^(max b1 b2 -1) - (-?v+2^(max b1 b2-1)-1) mod (2^max b1 b2) - 1" + assume "\ ?v\0" + hence "createSInt (max b1 b2) ?v = (ShowL\<^sub>i\<^sub>n\<^sub>t ?x)" unfolding createSInt_def by simp + moreover have "sub (TSInt b1) (TSInt b2) (ShowL\<^sub>i\<^sub>n\<^sub>t v1) (ShowL\<^sub>i\<^sub>n\<^sub>t v2) + = Some (createSInt (max b1 b2) ?v, TSInt (max b1 b2))" + using Read_ShowL_id sub_def olift.simps(1)[of "(-)" b1 b2] by simp + ultimately have "expr (MINUS e1 e2) env cd st g + = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TSInt (max b1 b2))),g)" using r1 r2 True by (simp add:Statements.solidity.expr.simps) + moreover have "expr (eupdate (MINUS e1 e2)) env cd st g + = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TSInt (max b1 b2))),g)" + proof - + from u_def have "eupdate (MINUS e1 e2) = E.INT (max b1 b2) ?x" using `\ ?v\0` by simp + moreover have "expr (E.INT (max b1 b2) ?x) env cd st g + = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TSInt (max b1 b2))),g)" + proof - + from `b1 \ vbits` `b2 \ vbits` have "max b1 b2 \ vbits" using vbits_max by simp + with True have "expr (E.INT (max b1 b2) ?x) env cd st g + = Normal ((KValue (createSInt (max b1 b2) ?x), Value (TSInt (max b1 b2))),g)" by (simp add:Statements.solidity.expr.simps createSInt_def) + moreover from `0 < b1` + have "?x \ - (2 ^ (max b1 b2 - 1))" using lower_bound2[of "max b1 b2" ?v] by simp + moreover from `b1 > 0` have "2^(max b1 b2 -1) > (0::nat)" by simp + hence "2^(max b1 b2 -1) - (-?v+2^(max b1 b2-1)-1) mod (2^max b1 b2) - 1 < 2 ^ (max b1 b2 - 1)" + by (simp add: algebra_simps flip: int_one_le_iff_zero_less) + moreover from `0 < b1` have "0 < max b1 b2" using max_def by simp + ultimately show ?thesis using createSInt_id[of ?x "max b1 b2"] by simp + qed + ultimately show ?thesis by simp + qed + ultimately show ?thesis by simp + qed + next + assume "\ b2 \ vbits" + with m i i2 show ?thesis by (simp add:Statements.solidity.expr.simps) + qed + next + case u: (UINT b2 v2) + with m.IH have expr2: "expr e2 env cd st g = expr (UINT b2 v2) env cd st g" by simp + then show ?thesis + proof (cases) + let ?v="v1 - v2" + assume "b2 \ vbits" + with expr2 True + have "expr e2 env cd st g = Normal ((KValue (createUInt b2 v2), Value (TUInt b2)),g)" by (simp add:Statements.solidity.expr.simps) + moreover from u `b2 \ vbits` + have "v2 < 2^b2" and "v2 \ 0" using update_bounds_uint by auto + moreover from `b2 \ vbits` have "0 < b2" by auto + ultimately have r2: "expr e2 env cd st g = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t v2), Value (TUInt b2)),g)" + using createUInt_id[of v2 b2] by simp + thus ?thesis + proof (cases) + assume "b2 vbits` `b2 \ vbits` have + u_def: "eupdate (MINUS e1 e2) = + (let v = v1 - v2 + in if 0 \ v + then E.INT b1 (- (2 ^ (b1 - 1)) + (v + 2 ^ (b1 - 1)) mod 2 ^ b1) + else E.INT b1 (2 ^ (b1 - 1) - (- v + 2 ^ (b1 - 1) - 1) mod 2 ^ b1 - 1))" + using i u eupdate.simps(11)[of e1 e2] by simp + show ?thesis + proof (cases) + let ?x="- (2 ^ (b1 - 1)) + (?v + 2 ^ (b1 - 1)) mod 2 ^ b1" + assume "?v\0" + hence "createSInt b1 ?v = (ShowL\<^sub>i\<^sub>n\<^sub>t ?x)" using `b2i\<^sub>n\<^sub>t v1) (ShowL\<^sub>i\<^sub>n\<^sub>t v2) + = Some (createSInt b1 ?v, TSInt b1)" + using Read_ShowL_id sub_def olift.simps(3)[of "(-)" b1 b2] `b2i\<^sub>n\<^sub>t ?x), Value (TSInt b1)),g)" using r1 r2 True by (simp add:Statements.solidity.expr.simps) + moreover have "expr (eupdate (MINUS e1 e2)) env cd st g + = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TSInt b1)),g)" + proof - + from u_def have "eupdate (MINUS e1 e2) = E.INT b1 ?x" using `?v\0` by simp + moreover have "expr (E.INT b1 ?x) env cd st g + = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TSInt b1)),g)" + proof - + from `b1 \ vbits` True have "expr (E.INT b1 ?x) env cd st g + = Normal ((KValue (createSInt b1 ?x), Value (TSInt b1)),g)" by (simp add:Statements.solidity.expr.simps) + moreover from `0 < b1` have "?x < 2 ^ (b1 - 1)" using upper_bound2 by simp + ultimately show ?thesis using createSInt_id[of ?x "b1"] `0 < b1` by simp + qed + ultimately show ?thesis by simp + qed + ultimately show ?thesis by simp + next + let ?x="2^(b1 -1) - (-?v+2^(b1-1)-1) mod (2^b1) - 1" + assume "\ ?v\0" + hence "createSInt b1 ?v = (ShowL\<^sub>i\<^sub>n\<^sub>t ?x)" unfolding createSInt_def by simp + moreover have "sub (TSInt b1) (TUInt b2) (ShowL\<^sub>i\<^sub>n\<^sub>t v1) (ShowL\<^sub>i\<^sub>n\<^sub>t v2) + = Some (createSInt b1 ?v, TSInt b1)" + using Read_ShowL_id sub_def olift.simps(3)[of "(-)" b1 b2] `b2i\<^sub>n\<^sub>t ?x), Value (TSInt b1)),g)" using r1 r2 True by (simp add:Statements.solidity.expr.simps) + moreover have "expr (eupdate (MINUS e1 e2)) env cd st g + = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TSInt b1)),g)" + proof - + from u_def have "eupdate (MINUS e1 e2) = E.INT b1 ?x" using `\ ?v\0` by simp + moreover have "expr (E.INT b1 ?x) env cd st g + = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TSInt b1)),g)" + proof - + from `b1 \ vbits` True have "expr (E.INT b1 ?x) env cd st g + = Normal ((KValue (createSInt b1 ?x), Value (TSInt b1)),g)" by (simp add:Statements.solidity.expr.simps) + moreover from `0 < b1` have "?x \ - (2 ^ (b1 - 1))" using upper_bound2 by simp + moreover have "2^(b1-1) - (-?v+2^(b1-1)-1) mod (2^b1) - 1 < 2 ^ (b1 - 1)" + by (simp add: algebra_simps flip: int_one_le_iff_zero_less) + ultimately show ?thesis using createSInt_id[of ?x b1] `0 < b1` by simp + qed + ultimately show ?thesis by simp + qed + ultimately show ?thesis by simp + qed + next + assume "\ b2 < b1" + with i u have "eupdate (MINUS e1 e2) = MINUS (eupdate e1) (eupdate e2)" by simp + with m i show ?thesis by (simp add:Statements.solidity.expr.simps) + qed + next + assume "\ b2 \ vbits" + with m i u show ?thesis by (simp add:Statements.solidity.expr.simps) + qed + next + case (ADDRESS _) + with m i show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (BALANCE _) + with m i show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case THIS + with m i show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case SENDER + with m i show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case VALUE + with m i show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case TRUE + with m i show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case FALSE + with m i show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (LVAL _) + with m i show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (PLUS _ _) + with m i show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (MINUS _ _) + with m i show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (EQUAL _ _) + with m i show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (LESS _ _) + with m i show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (AND _ _) + with m i show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (OR _ _) + with m i show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (NOT _) + with m i show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (CALL x181 x182) + with m i show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (ECALL x191 x192 x193) + with m i show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case CONTRACTS + with m i show ?thesis by (simp add:Statements.solidity.expr.simps) + qed + next + assume "\ b1 \ vbits" + with m i show ?thesis by (simp add:Statements.solidity.expr.simps) + qed + next + case False + then show ?thesis using no_gas by simp + qed + next + case u: (UINT b1 v1) + with m.IH have expr1: "expr e1 env cd st g = expr (UINT b1 v1) env cd st g" by simp + then show ?thesis + proof (cases "g > 0") + case True + show ?thesis + proof (cases) + assume "b1 \ vbits" + with expr1 True + have "expr e1 env cd st g = Normal ((KValue (createUInt b1 v1), Value (TUInt b1)),g)" by (simp add:Statements.solidity.expr.simps) + moreover from u `b1 \ vbits` + have "v1 < 2^b1" and "v1 \ 0" using update_bounds_uint by auto + moreover from `b1 \ vbits` have "0 < b1" by auto + ultimately have r1: "expr e1 env cd st g = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t v1), Value (TUInt b1)),g)" + by (simp add:Statements.solidity.expr.simps createUInt_def) + thus ?thesis + proof (cases "eupdate e2") + case u2: (UINT b2 v2) + with m.IH have expr2: "expr e2 env cd st g = expr (UINT b2 v2) env cd st g" by simp + then show ?thesis + proof (cases) + let ?v="v1 - v2" + let ?x="?v mod 2 ^ max b1 b2" + assume "b2 \ vbits" + with expr2 True + have "expr e2 env cd st g = Normal ((KValue (createUInt b2 v2), Value (TUInt b2)),g)" by (simp add:Statements.solidity.expr.simps) + moreover from u2 `b2 \ vbits` + have "v2 < 2^b2" and "v2 \ 0" using update_bounds_uint by auto + moreover from `b2 \ vbits` have "0 < b2" by auto + ultimately have r2: "expr e2 env cd st g = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t v2), Value (TUInt b2)),g)" + by (simp add:Statements.solidity.expr.simps createUInt_def) + moreover have "sub (TUInt b1) (TUInt b2) (ShowL\<^sub>i\<^sub>n\<^sub>t v1) (ShowL\<^sub>i\<^sub>n\<^sub>t v2) + = Some (createUInt (max b1 b2) ?v, TUInt (max b1 b2))" + using Read_ShowL_id sub_def olift.simps(2)[of "(-)" b1 b2] by simp + ultimately have "expr (MINUS e1 e2) env cd st g + = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TUInt (max b1 b2))),g)" using r1 True by (simp add:Statements.solidity.expr.simps createUInt_def) + moreover have "expr (eupdate (MINUS e1 e2)) env cd st g + = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TUInt (max b1 b2))),g)" + proof - + from `b1 \ vbits` `b2 \ vbits` + have "eupdate (MINUS e1 e2) = UINT (max b1 b2) ?x" using u u2 by simp + moreover have "expr (UINT (max b1 b2) ?x) env cd st g + = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TUInt (max b1 b2))),g)" + proof - + from `b1 \ vbits` `b2 \ vbits` have "max b1 b2 \ vbits" using vbits_max by simp + with True have "expr (UINT (max b1 b2) ?x) env cd st g + = Normal ((KValue (createUInt (max b1 b2) ?x), Value (TUInt (max b1 b2))),g)" by (simp add:Statements.solidity.expr.simps) + moreover from `0 < b1` + have "?x < 2 ^ (max b1 b2)" by simp + moreover from `0 < b1` have "0 < max b1 b2" using max_def by simp + ultimately show ?thesis by (simp add:Statements.solidity.expr.simps createUInt_def) + qed + ultimately show ?thesis by simp + qed + ultimately show ?thesis by simp + next + assume "\ b2 \ vbits" + with m u u2 show ?thesis by (simp add:Statements.solidity.expr.simps) + qed + next + case i: (INT b2 v2) + with m.IH have expr2: "expr e2 env cd st g = expr (E.INT b2 v2) env cd st g" by simp + then show ?thesis + proof (cases) + let ?v="v1 - v2" + assume "b2 \ vbits" + with expr2 True + have "expr e2 env cd st g = Normal ((KValue (createSInt b2 v2), Value (TSInt b2)),g)" by (simp add:Statements.solidity.expr.simps) + moreover from i `b2 \ vbits` + have "v2 < 2^(b2-1)" and "v2 \ -(2^(b2-1))" using update_bounds_int by auto + moreover from `b2 \ vbits` have "0 < b2" by auto + ultimately have r2: "expr e2 env cd st g = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t v2), Value (TSInt b2)),g)" + using createSInt_id[of v2 b2] by simp + thus ?thesis + proof (cases) + assume "b1 vbits` `b2 \ vbits` have + u_def: "eupdate (MINUS e1 e2) = + (let v = v1 - v2 + in if 0 \ v + then E.INT b2 (- (2 ^ (b2 - 1)) + (v + 2 ^ (b2 - 1)) mod 2 ^ b2) + else E.INT b2 (2 ^ (b2 - 1) - (- v + 2 ^ (b2 - 1) - 1) mod 2 ^ b2 - 1))" + using u i by simp + show ?thesis + proof (cases) + let ?x="- (2 ^ (b2 - 1)) + (?v + 2 ^ (b2 - 1)) mod 2 ^ b2" + assume "?v\0" + hence "createSInt b2 ?v = (ShowL\<^sub>i\<^sub>n\<^sub>t ?x)" using `b1i\<^sub>n\<^sub>t v1) (ShowL\<^sub>i\<^sub>n\<^sub>t v2) + = Some (createSInt b2 ?v, TSInt b2)" + using Read_ShowL_id sub_def olift.simps(4)[of "(-)" b1 b2] `b1i\<^sub>n\<^sub>t ?x), Value (TSInt b2)),g)" using r1 r2 True by (simp add:Statements.solidity.expr.simps) + moreover have "expr (eupdate (MINUS e1 e2)) env cd st g + = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TSInt b2)),g)" + proof - + from u_def have "eupdate (MINUS e1 e2) = E.INT b2 ?x" using `?v\0` by simp + moreover have "expr (E.INT b2 ?x) env cd st g + = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TSInt b2)),g)" + proof - + from `b2 \ vbits` True have "expr (E.INT b2 ?x) env cd st g + = Normal ((KValue (createSInt b2 ?x), Value (TSInt b2)),g)" by (simp add:Statements.solidity.expr.simps) + moreover from `0 < b2` have "?x < 2 ^ (b2 - 1)" using upper_bound2 by simp + ultimately show ?thesis using createSInt_id[of ?x "b2"] `0 < b2` by simp + qed + ultimately show ?thesis by simp + qed + ultimately show ?thesis by simp + next + let ?x="2^(b2 -1) - (-?v+2^(b2-1)-1) mod (2^b2) - 1" + assume "\ ?v\0" + hence "createSInt b2 ?v = (ShowL\<^sub>i\<^sub>n\<^sub>t ?x)" unfolding createSInt_def by simp + moreover have "sub (TUInt b1) (TSInt b2) (ShowL\<^sub>i\<^sub>n\<^sub>t v1) (ShowL\<^sub>i\<^sub>n\<^sub>t v2) + = Some (createSInt b2 ?v, TSInt b2)" + using Read_ShowL_id sub_def olift.simps(4)[of "(-)" b1 b2] `b1i\<^sub>n\<^sub>t ?x), Value (TSInt b2)),g)" using r1 r2 True by (simp add:Statements.solidity.expr.simps) + moreover have "expr (eupdate (MINUS e1 e2)) env cd st g + = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TSInt b2)),g)" + proof - + from u_def have "eupdate (MINUS e1 e2) = E.INT b2 ?x" using `\ ?v\0` by simp + moreover have "expr (E.INT b2 ?x) env cd st g + = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TSInt b2)),g)" + proof - + from `b2 \ vbits` True have "expr (E.INT b2 ?x) env cd st g + = Normal ((KValue (createSInt b2 ?x), Value (TSInt b2)),g)" by (simp add:Statements.solidity.expr.simps) + moreover from `0 < b2` have "?x \ - (2 ^ (b2 - 1))" using upper_bound2 by simp + moreover have "2^(b2-1) - (-?v+2^(b2-1)-1) mod (2^b2) - 1 < 2 ^ (b2 - 1)" + by (simp add: algebra_simps flip: int_one_le_iff_zero_less) + ultimately show ?thesis using createSInt_id[of ?x b2] `0 < b2` by simp + qed + ultimately show ?thesis by simp + qed + ultimately show ?thesis by simp + qed + next + assume "\ b1 < b2" + with m u i show ?thesis by (simp add:Statements.solidity.expr.simps) + qed + next + assume "\ b2 \ vbits" + with m u i show ?thesis by (simp add:Statements.solidity.expr.simps) + qed + next + case (ADDRESS _) + with m u show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (BALANCE _) + with m u show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case THIS + with m u show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case SENDER + with m u show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case VALUE + with m u show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case TRUE + with m u show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case FALSE + with m u show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (LVAL _) + with m u show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (PLUS _ _) + with m u show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (MINUS _ _) + with m u show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (EQUAL _ _) + with m u show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (LESS _ _) + with m u show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (AND _ _) + with m u show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (OR _ _) + with m u show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (NOT _) + with m u show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (CALL x181 x182) + with m u show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (ECALL x191 x192 x193) + with m u show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case CONTRACTS + with m u show ?thesis by (simp add:Statements.solidity.expr.simps) + qed + next + assume "\ b1 \ vbits" + with m u show ?thesis by (simp add:Statements.solidity.expr.simps) + qed + next + case False + then show ?thesis using no_gas by simp + qed + next + case (ADDRESS x3) + with m show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (BALANCE x4) + with m show ?thesis using lift_eq[of e1 env cd st g "eupdate e1" e2 "eupdate e2"] by (auto simp add:Statements.solidity.expr.simps) + next + case THIS + with m show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case SENDER + with m show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case VALUE + with m show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case TRUE + with m show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case FALSE + with m show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (LVAL x7) + with m show ?thesis using lift_eq[of e1 env cd st g "eupdate e1" e2 "eupdate e2"] by (auto simp add:Statements.solidity.expr.simps) + next + case (PLUS x81 x82) + with m show ?thesis using lift_eq[of e1 env cd st g "eupdate e1" e2 "eupdate e2"] by (auto simp add:Statements.solidity.expr.simps) + next + case (MINUS x91 x92) + with m show ?thesis using lift_eq[of e1 env cd st g "eupdate e1" e2 "eupdate e2"] by (auto simp add:Statements.solidity.expr.simps) + next + case (EQUAL x101 x102) + with m show ?thesis using lift_eq[of e1 env cd st g "eupdate e1" e2 "eupdate e2"] by (auto simp add:Statements.solidity.expr.simps) + next + case (LESS x111 x112) + with m show ?thesis using lift_eq[of e1 env cd st g "eupdate e1" e2 "eupdate e2"] by (auto simp add:Statements.solidity.expr.simps) + next + case (AND x121 x122) + with m show ?thesis using lift_eq[of e1 env cd st g "eupdate e1" e2 "eupdate e2"] by (auto simp add:Statements.solidity.expr.simps) + next + case (OR x131 x132) + with m show ?thesis using lift_eq[of e1 env cd st g "eupdate e1" e2 "eupdate e2"] by (auto simp add:Statements.solidity.expr.simps) + next + case (NOT x131) + with m show ?thesis using lift_eq[of e1 env cd st g "eupdate e1" e2 "eupdate e2"] by (auto simp add:Statements.solidity.expr.simps) + next + case (CALL x181 x182) + with m show ?thesis using lift_eq[of e1 env cd st g "eupdate e1" e2 "eupdate e2"] by (simp add:Statements.solidity.expr.simps) + next + case (ECALL x191 x192 x193) + with m show ?thesis using lift_eq[of e1 env cd st g "eupdate e1" e2 "eupdate e2"] by (simp add:Statements.solidity.expr.simps) + next + case CONTRACTS + with m show ?thesis using lift_eq[of e1 env cd st g "eupdate e1" e2 "eupdate e2"] by (auto simp add:Statements.solidity.expr.simps) + qed +next + case e: (EQUAL e1 e2 g) + show ?case + proof (cases "eupdate e1") + case i: (INT b1 v1) + with e.IH have expr1: "expr e1 env cd st g = expr (E.INT b1 v1) env cd st g" by simp + then show ?thesis + proof (cases "g > 0") + case True + then show ?thesis + proof (cases) + assume "b1 \ vbits" + with expr1 True + have "expr e1 env cd st g = Normal ((KValue (createSInt b1 v1), Value (TSInt b1)),g)" by (simp add:Statements.solidity.expr.simps) + moreover from i `b1 \ vbits` + have "v1 < 2^(b1-1)" and "v1 \ -(2^(b1-1))" using update_bounds_int by auto + moreover from `b1 \ vbits` have "0 < b1" by auto + ultimately have r1: "expr e1 env cd st g = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t v1), Value (TSInt b1)),g)" + using createSInt_id[of v1 b1] by simp + thus ?thesis + proof (cases "eupdate e2") + case i2: (INT b2 v2) + with e.IH have expr2: "expr e2 env cd st g = expr (E.INT b2 v2) env cd st g" by simp + then show ?thesis + proof (cases) + assume "b2 \ vbits" + with expr2 True + have "expr e2 env cd st g = Normal ((KValue (createSInt b2 v2), Value (TSInt b2)),g)" by (simp add:Statements.solidity.expr.simps) + moreover from i2 `b2 \ vbits` + have "v2 < 2^(b2-1)" and "v2 \ -(2^(b2-1))" using update_bounds_int by auto + moreover from `b2 \ vbits` have "0 < b2" by auto + ultimately have r2: "expr e2 env cd st g = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t v2), Value (TSInt b2)),g)" + using createSInt_id[of v2 b2] by simp + with r1 True have "expr (EQUAL e1 e2) env cd st g = + Normal ((KValue (createBool ((ReadL\<^sub>i\<^sub>n\<^sub>t (ShowL\<^sub>i\<^sub>n\<^sub>t v1))=((ReadL\<^sub>i\<^sub>n\<^sub>t (ShowL\<^sub>i\<^sub>n\<^sub>t v2))))), Value TBool),g)" + using equal_def plift.simps(1)[of "(=)"] by (simp add:Statements.solidity.expr.simps) + hence "expr (EQUAL e1 e2) env cd st g = Normal ((KValue (createBool (v1=v2)), Value TBool),g)" + using Read_ShowL_id by simp + with `b1 \ vbits` `b2 \ vbits` True show ?thesis using i i2 by (simp add:Statements.solidity.expr.simps createBool_def) + next + assume "\ b2 \ vbits" + with e i i2 show ?thesis by (simp add:Statements.solidity.expr.simps) + qed + next + case u: (UINT b2 v2) + with e.IH have expr2: "expr e2 env cd st g = expr (UINT b2 v2) env cd st g" by simp + then show ?thesis + proof (cases) + assume "b2 \ vbits" + with expr2 True + have "expr e2 env cd st g = Normal ((KValue (createUInt b2 v2), Value (TUInt b2)),g)" by (simp add:Statements.solidity.expr.simps) + moreover from u `b2 \ vbits` + have "v2 < 2^b2" and "v2 \ 0" using update_bounds_uint by auto + moreover from `b2 \ vbits` have "0 < b2" by auto + ultimately have r2: "expr e2 env cd st g = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t v2), Value (TUInt b2)),g)" + using createUInt_id[of v2 b2] by simp + thus ?thesis + proof (cases) + assume "b2i\<^sub>n\<^sub>t (ShowL\<^sub>i\<^sub>n\<^sub>t v1))=((ReadL\<^sub>i\<^sub>n\<^sub>t (ShowL\<^sub>i\<^sub>n\<^sub>t v2))))), Value TBool),g)" + using equal_def plift.simps(3)[of "(=)"] by (simp add:Statements.solidity.expr.simps) + hence "expr (EQUAL e1 e2) env cd st g = Normal ((KValue (createBool (v1=v2)), Value TBool),g)" + using Read_ShowL_id by simp + with `b1 \ vbits` `b2 \ vbits` `b2 b2 < b1" + with i u have "eupdate (EQUAL e1 e2) = EQUAL (eupdate e1) (eupdate e2)" by simp + with e i show ?thesis by (simp add:Statements.solidity.expr.simps) + qed + next + assume "\ b2 \ vbits" + with e i u show ?thesis by (simp add:Statements.solidity.expr.simps) + qed + next + case (ADDRESS _) + with e i show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (BALANCE _) + with e i show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case THIS + with e i show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case SENDER + with e i show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case VALUE + with e i show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case TRUE + with e i show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case FALSE + with e i show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (LVAL _) + with e i show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (PLUS _ _) + with e i show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (MINUS _ _) + with e i show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (EQUAL _ _) + with e i show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (LESS _ _) + with e i show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (AND _ _) + with e i show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (OR _ _) + with e i show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (NOT _) + with e i show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (CALL x181 x182) + with e i show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (ECALL x191 x192 x193) + with e i show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case CONTRACTS + with e i show ?thesis by (simp add:Statements.solidity.expr.simps) + qed + next + assume "\ b1 \ vbits" + with e i show ?thesis by (simp add:Statements.solidity.expr.simps) + qed + next + case False + then show ?thesis using no_gas by simp + qed + next + case u: (UINT b1 v1) + with e.IH have expr1: "expr e1 env cd st g = expr (UINT b1 v1) env cd st g" by simp + then show ?thesis + proof (cases "g > 0") + case True + then show ?thesis + proof (cases) + assume "b1 \ vbits" + with expr1 True + have "expr e1 env cd st g = Normal ((KValue (createUInt b1 v1), Value (TUInt b1)),g)" by (simp add:Statements.solidity.expr.simps) + moreover from u `b1 \ vbits` + have "v1 < 2^b1" and "v1 \ 0" using update_bounds_uint by auto + moreover from `b1 \ vbits` have "0 < b1" by auto + ultimately have r1: "expr e1 env cd st g = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t v1), Value (TUInt b1)),g)" + by (simp add:Statements.solidity.expr.simps createUInt_def) + thus ?thesis + proof (cases "eupdate e2") + case u2: (UINT b2 v2) + with e.IH have expr2: "expr e2 env cd st g = expr (UINT b2 v2) env cd st g" by simp + then show ?thesis + proof (cases) + assume "b2 \ vbits" + with expr2 True + have "expr e2 env cd st g = Normal ((KValue (createUInt b2 v2), Value (TUInt b2)),g)" by (simp add:Statements.solidity.expr.simps) + moreover from u2 `b2 \ vbits` + have "v2 < 2^b2" and "v2 \ 0" using update_bounds_uint by auto + moreover from `b2 \ vbits` have "0 < b2" by auto + ultimately have r2: "expr e2 env cd st g = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t v2), Value (TUInt b2)),g)" + by (simp add:Statements.solidity.expr.simps createUInt_def) + with r1 True have "expr (EQUAL e1 e2) env cd st g = + Normal ((KValue (createBool ((ReadL\<^sub>i\<^sub>n\<^sub>t (ShowL\<^sub>i\<^sub>n\<^sub>t v1))=((ReadL\<^sub>i\<^sub>n\<^sub>t (ShowL\<^sub>i\<^sub>n\<^sub>t v2))))), Value TBool),g)" + using equal_def plift.simps(2)[of "(=)"] by (simp add:Statements.solidity.expr.simps createUInt_def) + hence "expr (EQUAL e1 e2) env cd st g = Normal ((KValue (createBool (v1=v2)), Value TBool),g)" + using Read_ShowL_id by simp + with `b1 \ vbits` `b2 \ vbits` show ?thesis using u u2 True by (simp add:Statements.solidity.expr.simps createBool_def) + next + assume "\ b2 \ vbits" + with e u u2 show ?thesis by (simp add:Statements.solidity.expr.simps) + qed + next + case i: (INT b2 v2) + with e.IH have expr2: "expr e2 env cd st g = expr (E.INT b2 v2) env cd st g" by simp + then show ?thesis + proof (cases) + let ?v="v1 + v2" + assume "b2 \ vbits" + with expr2 True + have "expr e2 env cd st g = Normal ((KValue (createSInt b2 v2), Value (TSInt b2)),g)" by (simp add:Statements.solidity.expr.simps) + moreover from i `b2 \ vbits` + have "v2 < 2^(b2-1)" and "v2 \ -(2^(b2-1))" using update_bounds_int by auto + moreover from `b2 \ vbits` have "0 < b2" by auto + ultimately have r2: "expr e2 env cd st g = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t v2), Value (TSInt b2)),g)" + using createSInt_id[of v2 b2] by simp + thus ?thesis + proof (cases) + assume "b1i\<^sub>n\<^sub>t (ShowL\<^sub>i\<^sub>n\<^sub>t v1))=((ReadL\<^sub>i\<^sub>n\<^sub>t (ShowL\<^sub>i\<^sub>n\<^sub>t v2))))), Value TBool),g)" + using equal_def plift.simps(4)[of "(=)"] by (simp add:Statements.solidity.expr.simps) + hence "expr (EQUAL e1 e2) env cd st g = Normal ((KValue (createBool (v1=v2)), Value TBool),g)" + using Read_ShowL_id by simp + with `b1 \ vbits` `b2 \ vbits` `b1 b1 < b2" + with e u i show ?thesis by (simp add:Statements.solidity.expr.simps) + qed + next + assume "\ b2 \ vbits" + with e u i show ?thesis by (simp add:Statements.solidity.expr.simps) + qed + next + case (ADDRESS _) + with e u show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (BALANCE _) + with e u show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case THIS + with e u show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case SENDER + with e u show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case VALUE + with e u show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case TRUE + with e u show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case FALSE + with e u show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (LVAL _) + with e u show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (PLUS _ _) + with e u show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (MINUS _ _) + with e u show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (EQUAL _ _) + with e u show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (LESS _ _) + with e u show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (AND _ _) + with e u show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (OR _ _) + with e u show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (NOT _) + with e u show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (CALL x181 x182) + with e u show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (ECALL x191 x192 x193) + with e u show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case CONTRACTS + with e u show ?thesis by (simp add:Statements.solidity.expr.simps) + qed + next + assume "\ b1 \ vbits" + with e u show ?thesis by (simp add:Statements.solidity.expr.simps) + qed + next + case False + then show ?thesis using no_gas by simp + qed + next + case (ADDRESS x3) + with e show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (BALANCE x4) + with e show ?thesis using lift_eq[of e1 env cd st g "eupdate e1" e2 "eupdate e2"] by (auto simp add:Statements.solidity.expr.simps) + next + case THIS + with e show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case SENDER + with e show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case VALUE + with e show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case TRUE + with e show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case FALSE + with e show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (LVAL x7) + with e show ?thesis using lift_eq[of e1 env cd st g "eupdate e1" e2 "eupdate e2"] by (auto simp add:Statements.solidity.expr.simps) + next + case (PLUS x81 x82) + with e show ?thesis using lift_eq[of e1 env cd st g "eupdate e1" e2 "eupdate e2"] by (auto simp add:Statements.solidity.expr.simps) + next + case (MINUS x91 x92) + with e show ?thesis using lift_eq[of e1 env cd st g "eupdate e1" e2 "eupdate e2"] by (auto simp add:Statements.solidity.expr.simps) + next + case (EQUAL x101 x102) + with e show ?thesis using lift_eq[of e1 env cd st g "eupdate e1" e2 "eupdate e2"] by (auto simp add:Statements.solidity.expr.simps) + next + case (LESS x111 x112) + with e show ?thesis using lift_eq[of e1 env cd st g "eupdate e1" e2 "eupdate e2"] by (auto simp add:Statements.solidity.expr.simps) + next + case (AND x121 x122) + with e show ?thesis using lift_eq[of e1 env cd st g "eupdate e1" e2 "eupdate e2"] by (auto simp add:Statements.solidity.expr.simps) + next + case (OR x131 x132) + with e show ?thesis using lift_eq[of e1 env cd st g "eupdate e1" e2 "eupdate e2"] by (auto simp add:Statements.solidity.expr.simps) + next + case (NOT x131) + with e show ?thesis using lift_eq[of e1 env cd st g "eupdate e1" e2 "eupdate e2"] by (auto simp add:Statements.solidity.expr.simps) + next + case (CALL x181 x182) + with e show ?thesis using lift_eq[of e1 env cd st g "eupdate e1" e2 "eupdate e2"] by (simp add:Statements.solidity.expr.simps) + next + case (ECALL x191 x192 x193) + with e show ?thesis using lift_eq[of e1 env cd st g "eupdate e1" e2 "eupdate e2"] by (simp add:Statements.solidity.expr.simps) + next + case CONTRACTS + with e show ?thesis using lift_eq[of e1 env cd st g "eupdate e1" e2 "eupdate e2"] by (auto simp add:Statements.solidity.expr.simps) + qed +next + case l: (LESS e1 e2) + show ?case + proof (cases "eupdate e1") + case i: (INT b1 v1) + with l.IH have expr1: "expr e1 env cd st g = expr (E.INT b1 v1) env cd st g" by (simp add:Statements.solidity.expr.simps) + then show ?thesis + proof (cases "g > 0") + case True + then show ?thesis + proof (cases) + assume "b1 \ vbits" + with expr1 True + have "expr e1 env cd st g =Normal ((KValue (createSInt b1 v1), Value (TSInt b1)),g)" by (simp add:Statements.solidity.expr.simps) + moreover from i `b1 \ vbits` + have "v1 < 2^(b1-1)" and "v1 \ -(2^(b1-1))" using update_bounds_int by auto + moreover from `b1 \ vbits` have "0 < b1" by auto + ultimately have r1: "expr e1 env cd st g = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t v1), Value (TSInt b1)),g)" + using createSInt_id[of v1 b1] by (simp add:Statements.solidity.expr.simps) + thus ?thesis + proof (cases "eupdate e2") + case i2: (INT b2 v2) + with l.IH have expr2: "expr e2 env cd st g = expr (E.INT b2 v2) env cd st g" by simp + then show ?thesis + proof (cases) + assume "b2 \ vbits" + with expr2 True + have "expr e2 env cd st g = Normal ((KValue (createSInt b2 v2), Value (TSInt b2)),g)" by (simp add:Statements.solidity.expr.simps) + moreover from i2 `b2 \ vbits` + have "v2 < 2^(b2-1)" and "v2 \ -(2^(b2-1))" using update_bounds_int by auto + moreover from `b2 \ vbits` have "0 < b2" by auto + ultimately have r2: "expr e2 env cd st g = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t v2), Value (TSInt b2)),g)" + using createSInt_id[of v2 b2] by simp + with r1 True have "expr (LESS e1 e2) env cd st g = + Normal ((KValue (createBool ((ReadL\<^sub>i\<^sub>n\<^sub>t (ShowL\<^sub>i\<^sub>n\<^sub>t v1))<((ReadL\<^sub>i\<^sub>n\<^sub>t (ShowL\<^sub>i\<^sub>n\<^sub>t v2))))), Value TBool),g)" + using less_def plift.simps(1)[of "(<)"] by (simp add:Statements.solidity.expr.simps) + hence "expr (LESS e1 e2) env cd st g = Normal ((KValue (createBool (v1 vbits` `b2 \ vbits` show ?thesis using i i2 True by (simp add:Statements.solidity.expr.simps createBool_def) + next + assume "\ b2 \ vbits" + with l i i2 show ?thesis by (simp add:Statements.solidity.expr.simps) + qed + next + case u: (UINT b2 v2) + with l.IH have expr2: "expr e2 env cd st g = expr (UINT b2 v2) env cd st g" by simp + then show ?thesis + proof (cases) + assume "b2 \ vbits" + with expr2 True + have "expr e2 env cd st g = Normal ((KValue (createUInt b2 v2), Value (TUInt b2)),g)" by (simp add:Statements.solidity.expr.simps) + moreover from u `b2 \ vbits` + have "v2 < 2^b2" and "v2 \ 0" using update_bounds_uint by auto + moreover from `b2 \ vbits` have "0 < b2" by auto + ultimately have r2: "expr e2 env cd st g = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t v2), Value (TUInt b2)),g)" + using createUInt_id[of v2 b2] by simp + thus ?thesis + proof (cases) + assume "b2i\<^sub>n\<^sub>t (ShowL\<^sub>i\<^sub>n\<^sub>t v1))<((ReadL\<^sub>i\<^sub>n\<^sub>t (ShowL\<^sub>i\<^sub>n\<^sub>t v2))))), Value TBool),g)" + using less_def plift.simps(3)[of "(<)"] by (simp add:Statements.solidity.expr.simps) + hence "expr (LESS e1 e2) env cd st g = Normal ((KValue (createBool (v1 vbits` `b2 \ vbits` `b2 b2 < b1" + with i u have "eupdate (LESS e1 e2) = LESS (eupdate e1) (eupdate e2)" by simp + with l i show ?thesis by (simp add:Statements.solidity.expr.simps) + qed + next + assume "\ b2 \ vbits" + with l i u show ?thesis by (simp add:Statements.solidity.expr.simps) + qed + next + case (ADDRESS _) + with l i show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (BALANCE _) + with l i show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case THIS + with l i show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case SENDER + with l i show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case VALUE + with l i show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case TRUE + with l i show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case FALSE + with l i show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (LVAL _) + with l i show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (PLUS _ _) + with l i show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (MINUS _ _) + with l i show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (EQUAL _ _) + with l i show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (LESS _ _) + with l i show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (AND _ _) + with l i show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (OR _ _) + with l i show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (NOT _) + with l i show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (CALL x181 x182) + with l i show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (ECALL x191 x192 x193) + with l i show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case CONTRACTS + with l i show ?thesis by (simp add:Statements.solidity.expr.simps) + qed + next + assume "\ b1 \ vbits" + with l i show ?thesis by (simp add:Statements.solidity.expr.simps) + qed + next + case False + then show ?thesis using no_gas by (simp add:Statements.solidity.expr.simps) + qed + next + case u: (UINT b1 v1) + with l.IH have expr1: "expr e1 env cd st g = expr (UINT b1 v1) env cd st g" by simp + then show ?thesis + proof (cases "g > 0") + case True + then show ?thesis + proof (cases) + assume "b1 \ vbits" + with expr1 True + have "expr e1 env cd st g = Normal ((KValue (createUInt b1 v1), Value (TUInt b1)),g)" by (simp add:Statements.solidity.expr.simps) + moreover from u `b1 \ vbits` + have "v1 < 2^b1" and "v1 \ 0" using update_bounds_uint by auto + moreover from `b1 \ vbits` have "0 < b1" by auto + ultimately have r1: "expr e1 env cd st g = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t v1), Value (TUInt b1)),g)" + by (simp add:Statements.solidity.expr.simps createUInt_def) + thus ?thesis + proof (cases "eupdate e2") + case u2: (UINT b2 v2) + with l.IH have expr2: "expr e2 env cd st g = expr (UINT b2 v2) env cd st g" by simp + then show ?thesis + proof (cases) + assume "b2 \ vbits" + with expr2 True + have "expr e2 env cd st g = Normal ((KValue (createUInt b2 v2), Value (TUInt b2)),g)" by (simp add:Statements.solidity.expr.simps) + moreover from u2 `b2 \ vbits` + have "v2 < 2^b2" and "v2 \ 0" using update_bounds_uint by auto + moreover from `b2 \ vbits` have "0 < b2" by auto + ultimately have r2: "expr e2 env cd st g = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t v2), Value (TUInt b2)),g)" + by (simp add:Statements.solidity.expr.simps createUInt_def) + with r1 True have "expr (LESS e1 e2) env cd st g = Normal ((KValue (createBool ((ReadL\<^sub>i\<^sub>n\<^sub>t (ShowL\<^sub>i\<^sub>n\<^sub>t v1))<((ReadL\<^sub>i\<^sub>n\<^sub>t (ShowL\<^sub>i\<^sub>n\<^sub>t v2))))), Value TBool),g)" using less_def plift.simps(2)[of "(<)"] by (simp add:Statements.solidity.expr.simps) + hence "expr (LESS e1 e2) env cd st g = Normal ((KValue (createBool (v1 vbits` `b2 \ vbits` show ?thesis using u u2 True by (simp add:Statements.solidity.expr.simps createBool_def) + next + assume "\ b2 \ vbits" + with l u u2 show ?thesis by (simp add:Statements.solidity.expr.simps) + qed + next + case i: (INT b2 v2) + with l.IH have expr2: "expr e2 env cd st g = expr (E.INT b2 v2) env cd st g" by simp + then show ?thesis + proof (cases) + let ?v="v1 + v2" + assume "b2 \ vbits" + with expr2 True + have "expr e2 env cd st g = Normal ((KValue (createSInt b2 v2), Value (TSInt b2)),g)" by (simp add:Statements.solidity.expr.simps) + moreover from i `b2 \ vbits` + have "v2 < 2^(b2-1)" and "v2 \ -(2^(b2-1))" using update_bounds_int by auto + moreover from `b2 \ vbits` have "0 < b2" by auto + ultimately have r2: "expr e2 env cd st g = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t v2), Value (TSInt b2)),g)" + using createSInt_id[of v2 b2] by simp + thus ?thesis + proof (cases) + assume "b1i\<^sub>n\<^sub>t (ShowL\<^sub>i\<^sub>n\<^sub>t v1))<((ReadL\<^sub>i\<^sub>n\<^sub>t (ShowL\<^sub>i\<^sub>n\<^sub>t v2))))), Value TBool),g)" + using less_def plift.simps(4)[of "(<)"] by (simp add:Statements.solidity.expr.simps) + hence "expr (LESS e1 e2) env cd st g = Normal ((KValue (createBool (v1 vbits` `b2 \ vbits` `b1 b1 < b2" + with l u i show ?thesis by (simp add:Statements.solidity.expr.simps) + qed + next + assume "\ b2 \ vbits" + with l u i show ?thesis by (simp add:Statements.solidity.expr.simps) + qed + next + case (ADDRESS _) + with l u show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (BALANCE _) + with l u show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case THIS + with l u show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case SENDER + with l u show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case VALUE + with l u show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case TRUE + with l u show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case FALSE + with l u show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (LVAL _) + with l u show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (PLUS _ _) + with l u show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (MINUS _ _) + with l u show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (EQUAL _ _) + with l u show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (LESS _ _) + with l u show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (AND _ _) + with l u show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (OR _ _) + with l u show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (NOT _) + with l u show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (CALL x181 x182) + with l u show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (ECALL x191 x192 x193) + with l u show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case CONTRACTS + with l u show ?thesis by (simp add:Statements.solidity.expr.simps) + qed + next + assume "\ b1 \ vbits" + with l u show ?thesis by (simp add:Statements.solidity.expr.simps) + qed + next + case False + then show ?thesis using no_gas by simp + qed + next + case (ADDRESS x3) + with l show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (BALANCE x4) + with l show ?thesis using lift_eq[of e1 env cd st g "eupdate e1" e2 "eupdate e2"] by (auto simp add:Statements.solidity.expr.simps) + next + case THIS + with l show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case SENDER + with l show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case VALUE + with l show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case TRUE + with l show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case FALSE + with l show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (LVAL x7) + with l show ?thesis using lift_eq[of e1 env cd st g "eupdate e1" e2 "eupdate e2"] by (auto simp add:Statements.solidity.expr.simps) + next + case (PLUS x81 x82) + with l show ?thesis using lift_eq[of e1 env cd st g "eupdate e1" e2 "eupdate e2"] by (auto simp add:Statements.solidity.expr.simps) + next + case (MINUS x91 x92) + with l show ?thesis using lift_eq[of e1 env cd st g "eupdate e1" e2 "eupdate e2"] by (auto simp add:Statements.solidity.expr.simps) + next + case (EQUAL x101 x102) + with l show ?thesis using lift_eq[of e1 env cd st g "eupdate e1" e2 "eupdate e2"] by (auto simp add:Statements.solidity.expr.simps) + next + case (LESS x111 x112) + with l show ?thesis using lift_eq[of e1 env cd st g "eupdate e1" e2 "eupdate e2"] by (auto simp add:Statements.solidity.expr.simps) + next + case (AND x121 x122) + with l show ?thesis using lift_eq[of e1 env cd st g "eupdate e1" e2 "eupdate e2"] by (auto simp add:Statements.solidity.expr.simps) + next + case (OR x131 x132) + with l show ?thesis using lift_eq[of e1 env cd st g "eupdate e1" e2 "eupdate e2"] by (auto simp add:Statements.solidity.expr.simps) + next + case (NOT x131) + with l show ?thesis using lift_eq[of e1 env cd st g "eupdate e1" e2 "eupdate e2"] by (auto simp add:Statements.solidity.expr.simps) + next + case (CALL x181 x182) + with l show ?thesis using lift_eq[of e1 env cd st g "eupdate e1" e2 "eupdate e2"] by (simp add:Statements.solidity.expr.simps) + next + case (ECALL x191 x192 x193) + with l show ?thesis using lift_eq[of e1 env cd st g "eupdate e1" e2 "eupdate e2"] by (simp add:Statements.solidity.expr.simps) + next + case CONTRACTS + with l show ?thesis using lift_eq[of e1 env cd st g "eupdate e1" e2 "eupdate e2"] by (auto simp add:Statements.solidity.expr.simps) + qed +next + case a: (AND e1 e2) + show ?case + proof (cases "eupdate e1") + case (INT x11 x12) + with a show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (UINT x21 x22) + with a show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (ADDRESS x3) + with a show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (BALANCE x4) + with a show ?thesis using lift_eq[of e1 env cd st g "eupdate e1" e2 "eupdate e2"] by (auto simp add:Statements.solidity.expr.simps) + next + case THIS + with a show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case SENDER + with a show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case VALUE + with a show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case t: TRUE + show ?thesis + proof (cases "eupdate e2") + case (INT x11 x12) + with a t show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (UINT x21 x22) + with a t show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (ADDRESS x3) + with a t show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (BALANCE x4) + with a t show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case THIS + with a t show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case SENDER + with a t show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case VALUE + with a t show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case TRUE + with a t show ?thesis by (simp add:Statements.solidity.expr.simps vtand.simps) + next + case FALSE + with a t show ?thesis by (simp add:Statements.solidity.expr.simps vtand.simps) + next + case (LVAL x7) + with a t show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (PLUS x81 x82) + with a t show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (MINUS x91 x92) + with a t show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (EQUAL x101 x102) + with a t show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (LESS x111 x112) + with a t show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (AND x121 x122) + with a t show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (OR x131 x132) + with a t show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (NOT x131) + with a t show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (CALL x181 x182) + with a t show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (ECALL x191 x192 x193) + with a t show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case CONTRACTS + with a t show ?thesis by (simp add:Statements.solidity.expr.simps) + qed + next + case f: FALSE + show ?thesis + proof (cases "eupdate e2") + case (INT b v) + with a f show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (UINT b v) + with a f show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (ADDRESS x3) + with a f show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (BALANCE x4) + with a f show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case THIS + with a f show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case SENDER + with a f show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case VALUE + with a f show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case TRUE + with a f show ?thesis by (simp add:Statements.solidity.expr.simps vtand.simps) + next + case FALSE + with a f show ?thesis by (simp add:Statements.solidity.expr.simps vtand.simps) + next + case (LVAL x7) + with a f show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (PLUS x81 x82) + with a f show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (MINUS x91 x92) + with a f show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (EQUAL x101 x102) + with a f show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (LESS x111 x112) + with a f show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (AND x121 x122) + with a f show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (OR x131 x132) + with a f show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (NOT x131) + with a f show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (CALL x181 x182) + with a f show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (ECALL x191 x192 x193) + with a f show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case CONTRACTS + with a f show ?thesis by (simp add:Statements.solidity.expr.simps) + qed + next + case (LVAL x7) + with a show ?thesis using lift_eq[of e1 env cd st g "eupdate e1" e2 "eupdate e2"] by (auto simp add:Statements.solidity.expr.simps) + next + case p: (PLUS x81 x82) + with a show ?thesis using lift_eq[of e1 env cd st g "eupdate e1" e2 "eupdate e2"] by (auto simp add:Statements.solidity.expr.simps) + next + case (MINUS x91 x92) + with a show ?thesis using lift_eq[of e1 env cd st g "eupdate e1" e2 "eupdate e2"] by (auto simp add:Statements.solidity.expr.simps) + next + case (EQUAL x101 x102) + with a show ?thesis using lift_eq[of e1 env cd st g "eupdate e1" e2 "eupdate e2"] by (auto simp add:Statements.solidity.expr.simps) + next + case (LESS x111 x112) + with a show ?thesis using lift_eq[of e1 env cd st g "eupdate e1" e2 "eupdate e2"] by (auto simp add:Statements.solidity.expr.simps) + next + case (AND x121 x122) + with a show ?thesis using lift_eq[of e1 env cd st g "eupdate e1" e2 "eupdate e2"] by (auto simp add:Statements.solidity.expr.simps) + next + case (OR x131 x132) + with a show ?thesis using lift_eq[of e1 env cd st g "eupdate e1" e2 "eupdate e2"] by (auto simp add:Statements.solidity.expr.simps) + next + case (NOT x131) + with a show ?thesis using lift_eq[of e1 env cd st g "eupdate e1" e2 "eupdate e2"] by (auto simp add:Statements.solidity.expr.simps) + next + case (CALL x181 x182) + with a show ?thesis using lift_eq[of e1 env cd st g "eupdate e1" e2 "eupdate e2"] by (simp add:Statements.solidity.expr.simps) + next + case (ECALL x191 x192 x193) + with a show ?thesis using lift_eq[of e1 env cd st g "eupdate e1" e2 "eupdate e2"] by (simp add:Statements.solidity.expr.simps) + next + case CONTRACTS + with a show ?thesis using lift_eq[of e1 env cd st g "eupdate e1" e2 "eupdate e2"] by (auto simp add:Statements.solidity.expr.simps) + qed +next + case o: (OR e1 e2) + show ?case + proof (cases "eupdate e1") + case (INT x11 x12) + with o show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (UINT x21 x22) + with o show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (ADDRESS x3) + with o show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (BALANCE x4) + with o show ?thesis using lift_eq[of e1 env cd st g "eupdate e1" e2 "eupdate e2"] by (auto simp add:Statements.solidity.expr.simps) + next + case THIS + with o show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case SENDER + with o show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case VALUE + with o show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case t: TRUE + show ?thesis + proof (cases "eupdate e2") + case (INT x11 x12) + with o t show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (UINT x21 x22) + with o t show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (ADDRESS x3) + with o t show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (BALANCE x4) + with o t show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case THIS + with o t show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case SENDER + with o t show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case VALUE + with o t show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case TRUE + with o t show ?thesis by (simp add:Statements.solidity.expr.simps vtor.simps) + next + case FALSE + with o t show ?thesis by (simp add:Statements.solidity.expr.simps vtor.simps) + next + case (LVAL x7) + with o t show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (PLUS x81 x82) + with o t show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (MINUS x91 x92) + with o t show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (EQUAL x101 x102) + with o t show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (LESS x111 x112) + with o t show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (AND x121 x122) + with o t show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (OR x131 x132) + with o t show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (NOT x131) + with o t show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (CALL x181 x182) + with o t show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (ECALL x191 x192 x193) + with o t show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case CONTRACTS + with o t show ?thesis by (simp add:Statements.solidity.expr.simps) + qed + next + case f: FALSE + show ?thesis + proof (cases "eupdate e2") + case (INT b v) + with o f show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (UINT b v) + with o f show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (ADDRESS x3) + with o f show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (BALANCE x4) + with o f show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case THIS + with o f show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case SENDER + with o f show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case VALUE + with o f show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case TRUE + with o f show ?thesis by (simp add:Statements.solidity.expr.simps vtor.simps) + next + case FALSE + with o f show ?thesis by (simp add:Statements.solidity.expr.simps vtor.simps) + next + case (LVAL x7) + with o f show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (PLUS x81 x82) + with o f show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (MINUS x91 x92) + with o f show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (EQUAL x101 x102) + with o f show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (LESS x111 x112) + with o f show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (AND x121 x122) + with o f show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (OR x131 x132) + with o f show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (NOT x131) + with o f show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (CALL x181 x182) + with o f show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (ECALL x191 x192 x193) + with o f show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case CONTRACTS + with o f show ?thesis by (simp add:Statements.solidity.expr.simps) + qed + next + case (LVAL x7) + with o show ?thesis using lift_eq[of e1 env cd st g "eupdate e1" e2 "eupdate e2"] by (auto simp add:Statements.solidity.expr.simps) + next + case p: (PLUS x81 x82) + with o show ?thesis using lift_eq[of e1 env cd st g "eupdate e1" e2 "eupdate e2"] by (auto simp add:Statements.solidity.expr.simps) + next + case (MINUS x91 x92) + with o show ?thesis using lift_eq[of e1 env cd st g "eupdate e1" e2 "eupdate e2"] by (auto simp add:Statements.solidity.expr.simps) + next + case (EQUAL x101 x102) + with o show ?thesis using lift_eq[of e1 env cd st g "eupdate e1" e2 "eupdate e2"] by (auto simp add:Statements.solidity.expr.simps) + next + case (LESS x111 x112) + with o show ?thesis using lift_eq[of e1 env cd st g "eupdate e1" e2 "eupdate e2"] by (auto simp add:Statements.solidity.expr.simps) + next + case (AND x121 x122) + with o show ?thesis using lift_eq[of e1 env cd st g "eupdate e1" e2 "eupdate e2"] by (auto simp add:Statements.solidity.expr.simps) + next + case (OR x131 x132) + with o show ?thesis using lift_eq[of e1 env cd st g "eupdate e1" e2 "eupdate e2"] by (auto simp add:Statements.solidity.expr.simps) + next + case (NOT x131) + with o show ?thesis using lift_eq[of e1 env cd st g "eupdate e1" e2 "eupdate e2"] by (auto simp add:Statements.solidity.expr.simps) + next + case (CALL x181 x182) + with o show ?thesis using lift_eq[of e1 env cd st g "eupdate e1" e2 "eupdate e2"] by (simp add:Statements.solidity.expr.simps) + next + case (ECALL x191 x192 x193) + with o show ?thesis using lift_eq[of e1 env cd st g "eupdate e1" e2 "eupdate e2"] by (simp add:Statements.solidity.expr.simps) + next + case CONTRACTS + with o show ?thesis using lift_eq[of e1 env cd st g "eupdate e1" e2 "eupdate e2"] by (auto simp add:Statements.solidity.expr.simps) + qed +next + case o: (NOT e) + show ?case + proof (cases "eupdate e") + case (INT x11 x12) + with o show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (UINT x21 x22) + with o show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (ADDRESS x3) + with o show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (BALANCE x4) + with o show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case THIS + with o show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case SENDER + with o show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case VALUE + with o show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case t: TRUE + with o show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case f: FALSE + with o show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (LVAL x7) + with o show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case p: (PLUS x81 x82) + with o show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (MINUS x91 x92) + with o show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (EQUAL x101 x102) + with o show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (LESS x111 x112) + with o show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (AND x121 x122) + with o show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (OR x131 x132) + with o show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (NOT x131) + with o show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (CALL x181 x182) + with o show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case (ECALL x191 x192 x193) + with o show ?thesis by (simp add:Statements.solidity.expr.simps) + next + case CONTRACTS + with o show ?thesis by (simp add:Statements.solidity.expr.simps) + qed +next + case (CALL x181 x182) + show ?case by simp +next + case (ECALL x191 x192 x193 x194) + show ?case by simp +next + case CONTRACTS + show ?case by simp +qed + +end diff --git a/thys/Solidity/Contracts.thy b/thys/Solidity/Contracts.thy new file mode 100644 --- /dev/null +++ b/thys/Solidity/Contracts.thy @@ -0,0 +1,615 @@ +section \Contracts\ +theory Contracts + imports Environment +begin + +subsection \Syntax of Contracts\ + +datatype L = Id Identifier + | Ref Identifier "E list" +and E = INT nat int + | UINT nat int + | ADDRESS String.literal + | BALANCE E + | THIS + | SENDER + | VALUE + | TRUE + | FALSE + | LVAL L + | PLUS E E + | MINUS E E + | EQUAL E E + | LESS E E + | AND E E + | OR E E + | NOT E + | CALL Identifier "E list" + | ECALL E Identifier "E list" + | CONTRACTS + +datatype S = SKIP + | BLOCK "(Identifier \ Type) \ (E option)" S + | ASSIGN L E + | TRANSFER E E + | COMP S S + | ITE E S S + | WHILE E S + | INVOKE Identifier "E list" + | EXTERNAL E Identifier "E list" E + | NEW Identifier "E list" E + +abbreviation + "vbits\{8,16,24,32,40,48,56,64,72,80,88,96,104,112,120,128, + 136,144,152,160,168,176,184,192,200,208,216,224,232,240,248,256}" + +lemma vbits_max[simp]: + assumes "b1 \ vbits" + and "b2 \ vbits" + shows "(max b1 b2) \ vbits" +proof - + consider (b1) "max b1 b2 = b1" | (b2) "max b1 b2 = b2" by (metis max_def) + then show ?thesis + proof cases + case b1 + then show ?thesis using assms(1) by simp + next + case b2 + then show ?thesis using assms(2) by simp + qed +qed + +lemma vbits_ge_0: "(x::nat)\vbits \ x>0" by auto + +subsection \State\ + +type_synonym Gas = nat + +record State = + accounts :: Accounts + stack :: Stack + memory :: MemoryT + storage :: "Address \ StorageT" + gas :: Gas + +lemma all_gas_le: + assumes "gas x<(gas y::nat)" + and "\z. gas z < gas y \ P z \ Q z" + shows "\z. gas z \ gas x \ P z \ Q z" using assms by simp + +lemma all_gas_less: + assumes "\z. gas z < gas y \ P z" + and "gas x\(gas y::nat)" + shows "\z. gas z < gas x \ P z" using assms by simp + +definition incrementAccountContracts :: "Address \ State \ State" + where "incrementAccountContracts ad st = st \accounts := (accounts st)(ad := (accounts st ad)\contracts := Suc (contracts (accounts st ad))\)\" + +lemma incrementAccountContracts_type[simp]: + "type (accounts (incrementAccountContracts ad st) ad') = type (accounts st ad')" + using incrementAccountContracts_def by simp + +lemma incrementAccountContracts_bal[simp]: + "bal (accounts (incrementAccountContracts ad st) ad') = bal (accounts st ad')" + using incrementAccountContracts_def by simp + +lemma incrementAccountContracts_stack[simp]: + "stack (incrementAccountContracts ad st) = stack st" + using incrementAccountContracts_def by simp + +lemma incrementAccountContracts_memory[simp]: + "memory (incrementAccountContracts ad st) = memory st" + using incrementAccountContracts_def by simp + +lemma incrementAccountContracts_storage[simp]: + "storage (incrementAccountContracts ad st) = storage st" + using incrementAccountContracts_def by simp + +lemma incrementAccountContracts_gas[simp]: + "gas (incrementAccountContracts ad st) = gas st" + using incrementAccountContracts_def by simp + +lemma gas_induct: + assumes "\s. \s'. gas s' < gas s \ P s' \ P s" + shows "P s" using assms by (rule Nat.measure_induct[of "\s. gas s"]) + +definition emptyStorage :: "Address \ StorageT" +where + "emptyStorage _ = {$$}" + +abbreviation mystate::State + where "mystate \ \ + accounts = emptyAccount, + stack = emptyStore, + memory = emptyStore, + storage = emptyStorage, + gas = 0 + \" + +datatype Ex = Gas | Err + +subsection \Contracts\ + +text \ + A contract consists of methods, functions, and storage variables. + + A method is a triple consisting of + \<^item> A list of formal parameters + \<^item> A flag to signal external methods + \<^item> A statement + + + A function is a pair consisting of + \<^item> A list of formal parameters + \<^item> A flag to signal external functions + \<^item> An expression +\ +datatype(discs_sels) Member = Method "(Identifier \ Type) list \ bool \ S" + | Function "(Identifier \ Type) list \ bool \ E" + | Var STypes + +text \ + A procedure environment assigns a contract to an address. + A contract consists of + \<^item> An assignment of contract to identifiers + \<^item> A constructor + \<^item> A fallback statement which is executed after money is beeing transferred to the contract. + + \url{https://docs.soliditylang.org/en/v0.8.6/contracts.html#fallback-function} +\ + +type_synonym Contract = "(Identifier, Member) fmap \ ((Identifier \ Type) list \ S) \ S" + +type_synonym Environment\<^sub>P = "(Identifier, Contract) fmap" + +definition init::"(Identifier, Member) fmap \ Identifier \ Environment \ Environment" + where "init ct i e = (case fmlookup ct i of + Some (Var tp) \ updateEnvDup i (Storage tp) (Storeloc i) e + | _ \ e)" + +lemma init_s11[simp]: + assumes "fmlookup ct i = Some (Var tp)" + shows "init ct i e = updateEnvDup i (Storage tp) (Storeloc i) e" + using assms init_def by simp + +lemma init_s12[simp]: + assumes "i |\| fmdom (denvalue e)" + shows "init ct i e = e" +proof (cases "fmlookup ct i") + case None + then show ?thesis using init_def by simp +next + case (Some a) + then show ?thesis + proof (cases a) + case (Method _) + with Some show ?thesis using init_def by simp + next + case (Function _) + with Some show ?thesis using init_def by simp + next + case (Var tp) + with Some have "init ct i e = updateEnvDup i (Storage tp) (Storeloc i) e" using init_def by simp + moreover from assms have "updateEnvDup i (Storage tp) (Storeloc i) e = e" by auto + ultimately show ?thesis by simp + qed +qed + +lemma init_s13[simp]: + assumes "fmlookup ct i = Some (Var tp)" + and "\ i |\| fmdom (denvalue e)" + shows "init ct i e = updateEnv i (Storage tp) (Storeloc i) e" + using assms init_def by auto + +lemma init_s21[simp]: + assumes "fmlookup ct i = None" + shows "init ct i e = e" + using assms init_def by auto + +lemma init_s22[simp]: + assumes "fmlookup ct i = Some (Method m)" + shows "init ct i e = e" + using assms init_def by auto + +lemma init_s23[simp]: + assumes "fmlookup ct i = Some (Function f)" + shows "init ct i e = e" + using assms init_def by auto + +lemma init_commte: "comp_fun_commute (init ct)" +proof + fix x y + show "init ct y \ init ct x = init ct x \ init ct y" + proof + fix e + show "(init ct y \ init ct x) e = (init ct x \ init ct y) e" + proof (cases "fmlookup ct x") + case None + then show ?thesis by simp + next + case s1: (Some a) + then show ?thesis + proof (cases a) + case (Method _) + with s1 show ?thesis by simp + next + case (Function _) + with s1 show ?thesis by simp + next + case v1: (Var tp) + then show ?thesis + proof (cases "x |\| fmdom (denvalue e)") + case True + with s1 v1 have *: "init ct x e = e" by auto + then show ?thesis + proof (cases "fmlookup ct y") + case None + then show ?thesis by simp + next + case s2: (Some a) + then show ?thesis + proof (cases a) + case (Method _) + with s2 show ?thesis by simp + next + case (Function _) + with s2 show ?thesis by simp + next + case v2: (Var tp') + then show ?thesis + proof (cases "y |\| fmdom (denvalue e)") + case t1: True + with s1 v1 True s2 v2 show ?thesis by fastforce + next + define e' where "e' = updateEnv y (Storage tp') (Storeloc y) e" + case False + with s2 v2 have "init ct y e = e'" using e'_def by auto + with s1 v1 True e'_def * show ?thesis by auto + qed + qed + qed + next + define e' where "e' = updateEnv x (Storage tp) (Storeloc x) e" + case f1: False + with s1 v1 have *: "init ct x e = e'" using e'_def by auto + then show ?thesis + proof (cases "fmlookup ct y") + case None + then show ?thesis by simp + next + case s3: (Some a) + then show ?thesis + proof (cases a) + case (Method _) + with s3 show ?thesis by simp + next + case (Function _) + with s3 show ?thesis by simp + next + case v2: (Var tp') + then show ?thesis + proof (cases "y |\| fmdom (denvalue e)") + case t1: True + with e'_def have "y |\| fmdom (denvalue e')" by simp + with s1 s3 v1 f1 v2 show ?thesis using e'_def by fastforce + next + define f' where "f' = updateEnv y (Storage tp') (Storeloc y) e" + define e'' where "e'' = updateEnv y (Storage tp') (Storeloc y) e'" + case f2: False + with s3 v2 have **: "init ct y e = f'" using f'_def by auto + show ?thesis + proof (cases "y = x") + case True + with s3 v2 e'_def have "init ct y e' = e'" by simp + moreover from s3 v2 True f'_def have "init ct x f' = f'" by simp + ultimately show ?thesis using True by simp + next + define f'' where "f'' = updateEnv x (Storage tp) (Storeloc x) f'" + case f3: False + with f2 have "\ y |\| fmdom (denvalue e')" using e'_def by simp + with s3 v2 e''_def have "init ct y e' = e''" by auto + with * have "(init ct y \ init ct x) e = e''" by simp + moreover have "init ct x f' = f''" + proof - + from s1 v1 have "init ct x f' = updateEnvDup x (Storage tp) (Storeloc x) f'" by simp + moreover from f1 f3 have "x |\| fmdom (denvalue f')" using f'_def by simp + ultimately show ?thesis using f''_def by auto + qed + moreover from f''_def e''_def f'_def e'_def f3 have "Some f'' = Some e''" using env_reorder_neq by simp + ultimately show ?thesis using ** by simp + qed + qed + qed + qed + qed + qed + qed + qed +qed + +lemma init_address[simp]: + "address (init ct i e) = address e" +proof (cases "fmlookup ct i") + case None + then show ?thesis by simp +next + case (Some a) + show ?thesis + proof (cases a) + case (Method _) + with Some show ?thesis by simp + next + case (Function _) + with Some show ?thesis by simp + next + case (Var _) + with Some show ?thesis using updateEnvDup_address updateEnvDup_sender by simp + qed +qed + +lemma init_sender[simp]: +"sender (init ct i e) = sender e" +proof (cases "fmlookup ct i") + case None + then show ?thesis by simp +next + case (Some a) + show ?thesis + proof (cases a) + case (Method _) + with Some show ?thesis by simp + next + case (Function _) + with Some show ?thesis by simp + next + case (Var _) + with Some show ?thesis using updateEnvDup_sender by simp + qed +qed + +lemma init_svalue[simp]: +"svalue (init ct i e) = svalue e" +proof (cases "fmlookup ct i") + case None + then show ?thesis by simp +next + case (Some a) + show ?thesis + proof (cases a) + case (Method _) + with Some show ?thesis by simp + next + case (Function _) + with Some show ?thesis by simp + next + case (Var _) + with Some show ?thesis using updateEnvDup_svalue by simp + qed +qed + +lemma ffold_init_ad_same[rule_format]: "\e'. ffold (init ct) e xs = e' \ address e' = address e \ sender e' = sender e \ svalue e' = svalue e" +proof (induct xs) + case empty + then show ?case by (simp add: ffold_def) +next + case (insert x xs) + then have *: "ffold (init ct) e (finsert x xs) = + init ct x (ffold (init ct) e xs)" using FSet.comp_fun_commute.ffold_finsert[OF init_commte] by simp + show ?case + proof (rule allI[OF impI]) + fix e' assume **: "ffold (init ct) e (finsert x xs) = e'" + with * obtain e'' where ***: "ffold (init ct) e xs = e''" by simp + with insert have "address e'' = address e \ sender e'' = sender e \ svalue e'' = svalue e" by blast + with * ** *** show "address e' = address e \ sender e' = sender e \ svalue e' = svalue e" using init_address init_sender init_svalue by metis + qed +qed + +lemma ffold_init_ad[simp]: "address (ffold (init ct) e xs) = address e" + using ffold_init_ad_same by simp + +lemma ffold_init_sender[simp]: "sender (ffold (init ct) e xs) = sender e" + using ffold_init_ad_same by simp + +lemma ffold_init_dom: + "fmdom (denvalue (ffold (init ct) e xs)) |\| fmdom (denvalue e) |\| xs" +proof (induct "xs") + case empty + then show ?case + proof + fix x + assume "x |\| fmdom (denvalue (ffold (init ct) e {||}))" + moreover have "ffold (init ct) e {||} = e" using FSet.comp_fun_commute.ffold_empty[OF init_commte, of ct e] by simp + ultimately show "x |\| fmdom (denvalue e) |\| {||}" by simp + qed +next + case (insert x xs) + then have *: "ffold (init ct) e (finsert x xs) = + init ct x (ffold (init ct) e xs)" using FSet.comp_fun_commute.ffold_finsert[OF init_commte] by simp + + show ?case + proof + fix x' assume "x' |\| fmdom (denvalue (ffold (init ct) e (finsert x xs)))" + with * have **: "x' |\| fmdom (denvalue (init ct x (ffold (init ct) e xs)))" by simp + then consider "x' |\| fmdom (denvalue (ffold (init ct) e xs))" | "x'=x" + proof (cases "fmlookup ct x") + case None + then show ?thesis using that ** by simp + next + case (Some a) + then show ?thesis + proof (cases a) + case (Method _) + then show ?thesis using Some ** that by simp + next + case (Function _) + then show ?thesis using Some ** that by simp + next + case (Var x2) + show ?thesis + proof (cases "x=x'") + case True + then show ?thesis using that by simp + next + case False + then have "fmlookup (denvalue (updateEnvDup x (Storage x2) (Storeloc x) (ffold (init ct) e xs))) x' = fmlookup (denvalue (ffold (init ct) e xs)) x'" using updateEnvDup_dup by simp + moreover from ** Some Var have ***:"x' |\| fmdom (denvalue (updateEnvDup x (Storage x2) (Storeloc x) (ffold (init ct) e xs)))" by simp + ultimately have "x' |\| fmdom (denvalue (ffold (init ct) e xs))" by (simp add: fmlookup_dom_iff) + then show ?thesis using that by simp + qed + qed + qed + then show "x' |\| fmdom (denvalue e) |\| finsert x xs" + proof cases + case 1 + then show ?thesis using insert.hyps by auto + next + case 2 + then show ?thesis by simp + qed + qed +qed + +lemma ffold_init_fmap: + assumes "fmlookup ct i = Some (Var tp)" + and "i |\| fmdom (denvalue e)" + shows "i|\|xs \ fmlookup (denvalue (ffold (init ct) e xs)) i = Some (Storage tp, Storeloc i)" +proof (induct "xs") + case empty + then show ?case by simp +next + case (insert x xs) + then have *: "ffold (init ct) e (finsert x xs) = + init ct x (ffold (init ct) e xs)" using FSet.comp_fun_commute.ffold_finsert[OF init_commte] by simp + + from insert.prems consider (a) "i |\| xs" | (b) "\ i |\| xs \ i = x" by auto + then show "fmlookup (denvalue (ffold (init ct) e (finsert x xs))) i = Some (Storage tp, Storeloc i)" + proof cases + case a + with insert.hyps(2) have "fmlookup (denvalue (ffold (init ct) e xs)) i = Some (Storage tp, Storeloc i)" by simp + moreover have "fmlookup (denvalue (init ct x (ffold (init ct) e xs))) i = fmlookup (denvalue (ffold (init ct) e xs)) i" + proof (cases "fmlookup ct x") + case None + then show ?thesis by simp + next + case (Some a) + then show ?thesis + proof (cases a) + case (Method _) + with Some show ?thesis by simp + next + case (Function _) + with Some show ?thesis by simp + next + case (Var x2) + with Some have "init ct x (ffold (init ct) e xs) = updateEnvDup x (Storage x2) (Storeloc x) (ffold (init ct) e xs)" using init_def[of ct x "(ffold (init ct) e xs)"] by simp + moreover from insert a have "i\x" by auto + then have "fmlookup (denvalue (updateEnvDup x (Storage x2) (Storeloc x) (ffold (init ct) e xs))) i = fmlookup (denvalue (ffold (init ct) e xs)) i" using updateEnvDup_dup[of x i] by simp + ultimately show ?thesis by simp + qed + qed + ultimately show ?thesis using * by simp + next + case b + with assms(1) have "fmlookup ct x = Some (Var tp)" by simp + moreover from b assms(2) have "\ x |\| fmdom (denvalue (ffold (init ct) e xs))" using ffold_init_dom by auto + ultimately have "init ct x (ffold (init ct) e xs) = updateEnv x (Storage tp) (Storeloc x) (ffold (init ct) e xs)" by auto + with b * show ?thesis by simp + qed +qed + +lemma ffold_init_fmdom: + assumes "fmlookup ct i = Some (Var tp)" + and "i |\| fmdom (denvalue e)" + shows "fmlookup (denvalue (ffold (init ct) e (fmdom ct))) i = Some (Storage tp, Storeloc i)" +proof - + from assms(1) have "i |\| fmdom ct" by (rule Finite_Map.fmdomI) + then show ?thesis using ffold_init_fmap[OF assms] by simp +qed + +text\The following definition allows for a more fine-grained configuration of the + code generator. +\ +definition ffold_init::"(String.literal, Member) fmap \ Environment \ String.literal fset \ Environment" where + \ffold_init ct a c = ffold (init ct) a c\ +declare ffold_init_def [simp] + +lemma ffold_init_code [code]: + \ffold_init ct a c = fold (init ct) (remdups (sorted_list_of_set (fset c))) a\ + using comp_fun_commute_on.fold_set_fold_remdups ffold.rep_eq + ffold_init_def init_commte sorted_list_of_fset.rep_eq + sorted_list_of_fset_simps(1) + by (metis comp_fun_commute.comp_fun_commute comp_fun_commute_on.intro order_refl) + +lemma bind_case_stackvalue_cong [fundef_cong]: + assumes "x = x'" + and "\v. x = KValue v \ f v s = f' v s" + and "\p. x = KCDptr p \ g p s = g' p s" + and "\p. x = KMemptr p \ h p s = h' p s" + and "\p. x = KStoptr p \ i p s = i' p s" + shows "(case x of KValue v \ f v | KCDptr p \ g p | KMemptr p \ h p | KStoptr p \ i p) s + = (case x' of KValue v \ f' v | KCDptr p \ g' p | KMemptr p \ h' p | KStoptr p \ i' p) s" + using assms by (cases x, auto) + +lemma bind_case_type_cong [fundef_cong]: + assumes "x = x'" + and "\t. x = Value t \ f t s = f' t s" + and "\t. x = Calldata t \ g t s = g' t s" + and "\t. x = Memory t \ h t s = h' t s" + and "\t. x = Storage t \ i t s = i' t s" + shows "(case x of Value t \ f t | Calldata t \ g t | Memory t \ h t | Storage t \ i t ) s + = (case x' of Value t \ f' t | Calldata t \ g' t | Memory t \ h' t | Storage t \ i' t) s" + using assms by (cases x, auto) + +lemma bind_case_denvalue_cong [fundef_cong]: + assumes "x = x'" + and "\a. x = (Stackloc a) \ f a s = f' a s" + and "\a. x = (Storeloc a) \ g a s = g' a s" + shows "(case x of (Stackloc a) \ f a | (Storeloc a) \ g a) s + = (case x' of (Stackloc a) \ f' a | (Storeloc a) \ g' a) s" + using assms by (cases x, auto) + +lemma bind_case_mtypes_cong [fundef_cong]: + assumes "x = x'" + and "\a t. x = (MTArray a t) \ f a t s = f' a t s" + and "\p. x = (MTValue p) \ g p s = g' p s" + shows "(case x of (MTArray a t) \ f a t | (MTValue p) \ g p) s + = (case x' of (MTArray a t) \ f' a t | (MTValue p) \ g' p) s" + using assms by (cases x, auto) + +lemma bind_case_stypes_cong [fundef_cong]: + assumes "x = x'" + and "\a t. x = (STArray a t) \ f a t s = f' a t s" + and "\a t. x = (STMap a t) \ g a t s = g' a t s" + and "\p. x = (STValue p) \ h p s = h' p s" + shows "(case x of (STArray a t) \ f a t | (STMap a t) \ g a t | (STValue p) \ h p) s + = (case x' of (STArray a t) \ f' a t | (STMap a t) \ g' a t | (STValue p) \ h' p) s" + using assms by (cases x, auto) + +lemma bind_case_types_cong [fundef_cong]: + assumes "x = x'" + and "\a. x = (TSInt a) \ f a s = f' a s" + and "\a. x = (TUInt a) \ g a s = g' a s" + and "x = TBool \ h s = h' s" + and "x = TAddr \ i s = i' s" + shows "(case x of (TSInt a) \ f a | (TUInt a) \ g a | TBool \ h | TAddr \ i) s + = (case x' of (TSInt a) \ f' a | (TUInt a) \ g' a | TBool \ h' | TAddr \ i') s" + using assms by (cases x, auto) + +lemma bind_case_contract_cong [fundef_cong]: + assumes "x = x'" + and "\a. x = Method a \ f a s = f' a s" + and "\a. x = Function a \ g a s = g' a s" + and "\a. x = Var a \ h a s = h' a s" + shows "(case x of Method a \ f a | Function a \ g a | Var a \ h a) s + = (case x' of Method a \ f' a | Function a \ g' a | Var a \ h' a) s" + using assms by (cases x, auto) + +lemma bind_case_memoryvalue_cong [fundef_cong]: + assumes "x = x'" + and "\a. x = MValue a \ f a s = f' a s" + and "\a. x = MPointer a \ g a s = g' a s" + shows "(case x of (MValue a) \ f a | (MPointer a) \ g a) s + = (case x' of (MValue a) \ f' a | (MPointer a) \ g' a) s" + using assms by (cases x, auto) + +end diff --git a/thys/Solidity/Environment.thy b/thys/Solidity/Environment.thy --- a/thys/Solidity/Environment.thy +++ b/thys/Solidity/Environment.thy @@ -1,488 +1,407 @@ -section\Environment and State\ -theory Environment -imports Accounts Storage StateMonad -begin - -subsection \Environment\ - -datatype Type = Value Types - | Calldata MTypes - | Memory MTypes - | Storage STypes - -datatype Denvalue = Stackloc Location - | Storeloc Location - -type_synonym Identifier = String.literal - -record Environment = - address :: Address (*address of executing contract*) - sender :: Address (*address of calling contract*) - svalue :: Valuetype (*money send*) - denvalue :: "(Identifier, Type \ Denvalue) fmap" - -fun identifiers :: "Environment \ Identifier fset" - where "identifiers e = fmdom (denvalue e)" - -fun emptyEnv :: "Address \ Address \ Valuetype \ Environment" - where "emptyEnv a s v = \address = a, sender = s, svalue =v, denvalue = fmempty\" - -definition eempty :: "Environment" - where "eempty = emptyEnv (STR '''') (STR '''') (STR '''')" - -declare eempty_def [solidity_symbex] - -fun updateEnv :: "Identifier \ Type \ Denvalue \ Environment \ Environment" - where "updateEnv i t v e = e \ denvalue := fmupd i (t,v) (denvalue e) \" - -fun updateEnvOption :: "Identifier \ Type \ Denvalue \ Environment \ Environment option" - where "updateEnvOption i t v e = (case fmlookup (denvalue e) i of - Some _ \ None - | None \ Some (updateEnv i t v e))" - -lemma updateEnvOption_address: "updateEnvOption i t v e = Some e' \ address e = address e'" -by (auto split:option.split_asm) - -fun updateEnvDup :: "Identifier \ Type \ Denvalue \ Environment \ Environment" - where "updateEnvDup i t v e = (case fmlookup (denvalue e) i of - Some _ \ e - | None \ updateEnv i t v e)" - -lemma updateEnvDup_address[simp]: "address (updateEnvDup i t v e) = address e" - by (auto split:option.split) - -lemma updateEnvDup_sender[simp]: "sender (updateEnvDup i t v e) = sender e" - by (auto split:option.split) - -lemma updateEnvDup_svalue[simp]: "svalue (updateEnvDup i t v e) = svalue e" - by (auto split:option.split) - -lemma updateEnvDup_dup: - assumes "i\i'" shows "fmlookup (denvalue (updateEnvDup i t v e)) i' = fmlookup (denvalue e) i'" -proof (cases "fmlookup (denvalue e) i = None") - case True - then show ?thesis using assms by simp -next - case False - then obtain e' where "fmlookup (denvalue e) i = Some e'" by auto - then show ?thesis using assms by simp -qed - -lemma env_reorder_neq: - assumes "x\y" - shows "updateEnv x t1 v1 (updateEnv y t2 v2 e) = updateEnv y t2 v2 (updateEnv x t1 v1 e)" -proof - - have "address (updateEnv x t1 v1 (updateEnv y t2 v2 e)) = address (updateEnv y t2 v2 (updateEnv x t1 v1 e))" by simp - moreover from assms have "denvalue (updateEnv x t1 v1 (updateEnv y t2 v2 e)) = denvalue (updateEnv y t2 v2 (updateEnv x t1 v1 e))" using Finite_Map.fmupd_reorder_neq[of x y "(t1,v1)" "(t2,v2)"] by simp - ultimately show ?thesis by simp -qed - -lemma uEO_in: - assumes "i |\| fmdom (denvalue e)" - shows "updateEnvOption i t v e = None" - using assms by auto - -lemma uEO_n_In: - assumes "\ i |\| fmdom (denvalue e)" - shows "updateEnvOption i t v e = Some (updateEnv i t v e)" - using assms by auto - -fun astack :: "Identifier \ Type \ Stackvalue \ Stack * Environment \ Stack * Environment" - where "astack i t v (s, e) = (push v s, (updateEnv i t (Stackloc (ShowL\<^sub>n\<^sub>a\<^sub>t (toploc s))) e))" - -subsection \State\ - -type_synonym Gas = nat - -record State = - accounts :: Accounts - stack :: Stack - memory :: MemoryT - storage :: "(Address,StorageT) fmap" - gas :: Gas - -datatype Ex = Gas | Err - -fun append :: "Identifier \ Type \ Stackvalue - \ CalldataT \ Environment \ (CalldataT \ Environment, Ex, State) state_monad" -where - "append id0 tp v cd e st = - (let (k, e') = astack id0 tp v (stack st, e) - in do { - modify (\st. st \stack := k\); - return (cd, e') - }) st" - -subsection \Declarations\ - -text \This function is used to declare a new variable: decl id tp val copy cd mem cd' env st -\begin{description} - \item[id] is the name of the variable - \item[tp] is the type of the variable - \item[val] is an optional initialization parameter. If it is None, the types default value is taken. - \item[copy] is a flag to indicate whether memory should be copied (from mem parameter) or not (copying is required for example for external method calls). - \item[cd] is the original calldata which is used as a source - \item[mem] is the original memory which is used as a source - \item[cd'] is the new calldata - \item[env] is the new environment - \item[st] is the new state -\end{description}\ -fun decl :: "Identifier \ Type \ (Stackvalue * Type) option \ bool \ CalldataT \ MemoryT - \ CalldataT \ Environment \ (CalldataT \ Environment, Ex, State) state_monad" - where -(* Declaring stack variables *) - "decl i (Value t) None _ _ _ c env st = append i (Value t) (KValue (ival t)) c env st" -| "decl i (Value t) (Some (KValue v, Value t')) _ _ _ c env st = - (case convert t' t v of - Some (v', t'') \ append i (Value t'') (KValue v') c env - | None \ throw Err) st" -| "decl _ (Value _) (Some _) _ _ _ _ _ st = throw Err st" - -(* Declaring calldata variables *) -| "decl i (Calldata (MTArray x t)) (Some (KCDptr p, _)) True cd _ c env st = - (let l = ShowL\<^sub>n\<^sub>a\<^sub>t (toploc c); - (_, c') = allocate c - in (case cpm2m p l x t cd c' of - Some c'' \ append i (Calldata (MTArray x t)) (KCDptr l) c'' env - | None \ throw Err)) st" -| "decl i (Calldata (MTArray x t)) (Some (KMemptr p, _)) True _ mem c env st = - (let l = ShowL\<^sub>n\<^sub>a\<^sub>t (toploc c); - (_, c') = allocate c - in (case cpm2m p l x t mem c' of - Some c'' \ append i (Calldata (MTArray x t)) (KCDptr l) c'' env - | None \ throw Err)) st" -| "decl i (Calldata _) _ _ _ _ _ _ st = throw Err st" - -(* Declaring memory variables *) -| "decl i (Memory (MTArray x t)) None _ _ _ c env st = - (do { - m \ applyf (\st. memory st); - modify (\st. st \memory := minit x t m\); - append i (Memory (MTArray x t)) (KMemptr (ShowL\<^sub>n\<^sub>a\<^sub>t (toploc m))) c env - }) st" -| "decl i (Memory (MTArray x t)) (Some (KMemptr p, _)) True _ mem c env st = - (do { - m \ (applyf (\st. memory st)); - (case cpm2m p (ShowL\<^sub>n\<^sub>a\<^sub>t (toploc m)) x t mem (snd (allocate m)) of - Some m' \ - do { - modify (\st. st \memory := m'\); - append i (Memory (MTArray x t)) (KMemptr (ShowL\<^sub>n\<^sub>a\<^sub>t (toploc m))) c env - } - | None \ throw Err) - }) st" -| "decl i (Memory (MTArray x t)) (Some (KMemptr p, _)) False _ _ c env st = - append i (Memory (MTArray x t)) (KMemptr p) c env st" -| "decl i (Memory (MTArray x t)) (Some (KCDptr p, _)) _ cd _ c env st = - (do { - m \ (applyf (\st. memory st)); - (case cpm2m p (ShowL\<^sub>n\<^sub>a\<^sub>t (toploc m)) x t cd (snd (allocate m)) of - Some m' \ - do { - modify (\st. st \memory := m'\); - append i (Memory (MTArray x t)) (KMemptr (ShowL\<^sub>n\<^sub>a\<^sub>t (toploc m))) c env - } - | None \ throw Err) - }) st" -| "decl i (Memory (MTArray x t)) (Some (KStoptr p, Storage (STArray x' t'))) _ _ _ c env st = - (do { - s \ (applyf (\st. storage st)); - (case fmlookup s (address env) of - Some s' \ - (do { - m \ (applyf (\st. memory st)); - (case cps2m p (ShowL\<^sub>n\<^sub>a\<^sub>t (toploc m)) x' t' s' (snd (allocate m)) of - Some m'' \ - do { - modify (\st. st \memory := m''\); - append i (Memory (MTArray x t)) (KMemptr (ShowL\<^sub>n\<^sub>a\<^sub>t (toploc m))) c env - } - | None \ throw Err) - }) - | None \ throw Err) - }) st" -| "decl _ (Memory (MTArray _ _)) (Some _) _ _ _ _ _ st = throw Err st" -| "decl _ (Memory (MTValue _)) _ _ _ _ _ _ st = throw Err st" - -(* Declaring storage variables *) -| "decl _ (Storage (STArray _ _)) None _ _ _ _ _ st = throw Err st" (* Uninitialized storage arrays not allowed *) -| "decl i (Storage (STArray x t)) (Some (KStoptr p, _)) _ _ _ c env st = - append i (Storage (STArray x t)) (KStoptr p) c env st" -| "decl _ (Storage (STArray _ _)) (Some _) _ _ _ _ _ st = throw Err st" -| "decl _ (Storage (STMap _ _)) None _ _ _ _ _ st = throw Err st" (* Uninitialized storage maps not allowed *) -| "decl i (Storage (STMap t t')) (Some (KStoptr p, _)) _ _ _ c env st = - append i (Storage (STMap t t')) (KStoptr p) c env st" -| "decl _ (Storage (STMap _ _)) (Some _) _ _ _ _ _ st = throw Err st" -| "decl _ (Storage (STValue _)) _ _ _ _ _ _ st = throw Err st" - -(*TODO: refactor with decl.elims*) -lemma decl_gas_address: - assumes "decl a1 a2 a3 cp cd mem c env st = Normal ((l1', t1'), st1')" - shows "gas st1' = gas st \ address env = address t1' \ sender env = sender t1' \ svalue env = svalue t1'" -proof (cases a2) - case (Value t) - then show ?thesis - proof (cases a3) - case None - with Value show ?thesis using assms by auto - next - case (Some a) - show ?thesis - proof (cases a) - case (Pair a b) - then show ?thesis - proof (cases a) - case (KValue v) - then show ?thesis - proof (cases b) - case v2: (Value t') - show ?thesis - proof (cases "convert t' t v") - case None - with Some Pair KValue Value v2 show ?thesis using assms by simp - next - case s2: (Some a) - show ?thesis - proof (cases a) - case p2: (Pair a b) - with Some Pair KValue Value v2 s2 show ?thesis using assms by auto - qed - qed - next - case (Calldata x2) - with Some Pair KValue Value show ?thesis using assms by simp - next - case (Memory x3) - with Some Pair KValue Value show ?thesis using assms by simp - next - case (Storage x4) - with Some Pair KValue Value show ?thesis using assms by simp - qed - next - case (KCDptr x2) - with Some Pair Value show ?thesis using assms by simp - next - case (KMemptr x3) - with Some Pair Value show ?thesis using assms by simp - next - case (KStoptr x4) - with Some Pair Value show ?thesis using assms by simp - qed - qed - qed -next - case (Calldata x2) - then show ?thesis - proof (cases cp) - case True - then show ?thesis - proof (cases x2) - case (MTArray x t) - then show ?thesis - proof (cases a3) - case None - with Calldata show ?thesis using assms by simp - next - case (Some a) - show ?thesis - proof (cases a) - case (Pair a b) - then show ?thesis - proof (cases a) - case (KValue x1) - with Calldata Some Pair show ?thesis using assms by simp - next - case (KCDptr p) - define l where "l = ShowL\<^sub>n\<^sub>a\<^sub>t (toploc c)" - obtain c' where c_def: "\x. (x, c') = allocate c" by simp - show ?thesis - proof (cases "cpm2m p l x t cd c'") - case None - with Calldata MTArray Some Pair KCDptr l_def c_def True show ?thesis using assms by simp - next - case s2: (Some a) - with Calldata MTArray Some Pair KCDptr l_def c_def True show ?thesis using assms by auto - qed - next - case (KMemptr p) - define l where "l = ShowL\<^sub>n\<^sub>a\<^sub>t (toploc c)" - obtain c' where c_def: "\x. (x, c') = allocate c" by simp - show ?thesis - proof (cases "cpm2m p l x t mem c'") - case None - with Calldata MTArray Some Pair KMemptr l_def c_def True show ?thesis using assms by simp - next - case s2: (Some a) - with Calldata MTArray Some Pair KMemptr l_def c_def True show ?thesis using assms by auto - qed - next - case (KStoptr x4) - with Calldata Some Pair show ?thesis using assms by simp - qed - qed - qed - next - case (MTValue x2) - with Calldata show ?thesis using assms by simp - qed - next - case False - with Calldata show ?thesis using assms by simp - qed -next - case (Memory x3) - then show ?thesis - proof (cases x3) - case (MTArray x t) - then show ?thesis - proof (cases a3) - case None - with Memory MTArray None show ?thesis using assms by (auto simp add:Let_def) - next - case (Some a) - then show ?thesis - proof (cases a) - case (Pair a b) - then show ?thesis - proof (cases a) - case (KValue x1) - with Memory MTArray Some Pair show ?thesis using assms by simp - next - case (KCDptr p) - define m l where "m = memory st" and "l = ShowL\<^sub>n\<^sub>a\<^sub>t (toploc m)" - obtain m' where m'_def: "\x. (x, m') = allocate m" by simp - then show ?thesis - proof (cases "cpm2m p l x t cd m'") - case None - with Memory MTArray Some Pair KCDptr m_def l_def m'_def show ?thesis using assms by simp - next - case s2: (Some a) - with Memory MTArray Some Pair KCDptr m_def l_def m'_def show ?thesis using assms by auto - qed - next - case (KMemptr p) - then show ?thesis - proof (cases cp) - case True - define m l where "m = memory st" and "l = ShowL\<^sub>n\<^sub>a\<^sub>t (toploc m)" - obtain m' where m'_def: "\x. (x, m') = allocate m" by simp - then show ?thesis - proof (cases "cpm2m p l x t mem m'") - case None - with Memory MTArray Some Pair KMemptr True m_def l_def m'_def show ?thesis using assms by simp - next - case s2: (Some a) - with Memory MTArray Some Pair KMemptr True m_def l_def m'_def show ?thesis using assms by auto - qed - next - case False - with Memory MTArray Some Pair KMemptr show ?thesis using assms by auto - qed - next - case (KStoptr p) - then show ?thesis - proof (cases b) - case (Value x1) - with Memory MTArray Some Pair KStoptr show ?thesis using assms by simp - next - case (Calldata x2) - with Memory MTArray Some Pair KStoptr show ?thesis using assms by simp - next - case m2: (Memory x3) - with Memory MTArray Some Pair KStoptr show ?thesis using assms by simp - next - case (Storage x4) - then show ?thesis - proof (cases x4) - case (STArray x' t') - define m l where "m = memory st" and "l = ShowL\<^sub>n\<^sub>a\<^sub>t (toploc m)" - obtain m' where m'_def: "\x. (x, m') = allocate m" by simp - from assms(1) Memory MTArray Some Pair KStoptr Storage STArray m_def l_def m'_def - obtain s where *: "fmlookup (storage st) (address env) = Some s" using Let_def by (auto simp add: Let_def split:option.split_asm) - then show ?thesis - proof (cases "cps2m p l x' t' s m'") - case None - with Memory MTArray Some Pair KStoptr Storage STArray m_def l_def m'_def * show ?thesis using assms by simp - next - case s2: (Some a) - with Memory MTArray Some Pair KStoptr Storage STArray m_def l_def m'_def * show ?thesis using assms by auto - qed - next - case (STMap x21 x22) - with Memory MTArray Some Pair KStoptr Storage show ?thesis using assms by simp - next - case (STValue x3) - with Memory MTArray Some Pair KStoptr Storage show ?thesis using assms by simp - qed - qed - qed - qed - qed - next - case (MTValue x2) - with Memory show ?thesis using assms by simp - qed -next - case (Storage x4) - then show ?thesis - proof (cases x4) - case (STArray x t) - then show ?thesis - proof (cases a3) - case None - with Storage STArray show ?thesis using assms by simp - next - case (Some a) - then show ?thesis - proof (cases a) - case (Pair a b) - then show ?thesis - proof (cases a) - case (KValue x1) - with Storage STArray Some Pair show ?thesis using assms by simp - next - case (KCDptr x2) - with Storage STArray Some Pair show ?thesis using assms by simp - next - case (KMemptr x3) - with Storage STArray Some Pair show ?thesis using assms by simp - next - case (KStoptr x4) - with Storage STArray Some Pair show ?thesis using assms by auto - qed - qed - qed - next - case (STMap t t') - then show ?thesis - proof (cases a3) - case None - with Storage STMap show ?thesis using assms by simp - next - case (Some a) - then show ?thesis - proof (cases a) - case (Pair a b) - then show ?thesis - proof (cases a) - case (KValue x1) - with Storage STMap Some Pair show ?thesis using assms by simp - next - case (KCDptr x2) - with Storage STMap Some Pair show ?thesis using assms by simp - next - case (KMemptr x3) - with Storage STMap Some Pair show ?thesis using assms by simp - next - case (KStoptr x4) - with Storage STMap Some Pair show ?thesis using assms by auto - qed - qed - qed - next - case (STValue x3) - with Storage show ?thesis using assms by simp - qed -qed - -end +section\Environment and State\ +theory Environment +imports Accounts Storage StateMonad +begin + +subsection \Environment\ + +datatype Type = Value Types + | Calldata MTypes + | Memory MTypes + | Storage STypes + +datatype Denvalue = Stackloc Location + | Storeloc Location + +record Environment = + address :: Address (*address of executing contract*) + contract :: Identifier + sender :: Address (*address of calling contract*) + svalue :: Valuetype (*money send*) + denvalue :: "(Identifier, Type \ Denvalue) fmap" + +fun identifiers :: "Environment \ Identifier fset" + where "identifiers e = fmdom (denvalue e)" + +definition emptyEnv :: "Address \ Identifier \ Address \ Valuetype \ Environment" + where "emptyEnv a c s v = \address = a, contract = c, sender = s, svalue = v, denvalue = fmempty\" + +lemma emptyEnv_address[simp]: + "address (emptyEnv a c s v) = a" + unfolding emptyEnv_def by simp + +lemma emptyEnv_members[simp]: + "contract (emptyEnv a c s v) = c" + unfolding emptyEnv_def by simp + +lemma emptyEnv_sender[simp]: + "sender (emptyEnv a c s v) = s" + unfolding emptyEnv_def by simp + +lemma emptyEnv_svalue[simp]: + "svalue (emptyEnv a c s v) = v" + unfolding emptyEnv_def by simp + +lemma emptyEnv_denvalue[simp]: + "denvalue (emptyEnv a c s v) = {$$}" + unfolding emptyEnv_def by simp + +definition eempty :: "Environment" + where "eempty = emptyEnv (STR '''') (STR '''') (STR '''') (STR '''')" + +declare eempty_def [solidity_symbex] + +fun updateEnv :: "Identifier \ Type \ Denvalue \ Environment \ Environment" + where "updateEnv i t v e = e \ denvalue := fmupd i (t,v) (denvalue e) \" + +fun updateEnvOption :: "Identifier \ Type \ Denvalue \ Environment \ Environment option" + where "updateEnvOption i t v e = (case fmlookup (denvalue e) i of + Some _ \ None + | None \ Some (updateEnv i t v e))" + +lemma updateEnvOption_address: "updateEnvOption i t v e = Some e' \ address e = address e'" +by (auto split:option.split_asm) + +fun updateEnvDup :: "Identifier \ Type \ Denvalue \ Environment \ Environment" + where "updateEnvDup i t v e = (case fmlookup (denvalue e) i of + Some _ \ e + | None \ updateEnv i t v e)" + +lemma updateEnvDup_address[simp]: "address (updateEnvDup i t v e) = address e" + by (auto split:option.split) + +lemma updateEnvDup_sender[simp]: "sender (updateEnvDup i t v e) = sender e" + by (auto split:option.split) + +lemma updateEnvDup_svalue[simp]: "svalue (updateEnvDup i t v e) = svalue e" + by (auto split:option.split) + +lemma updateEnvDup_dup: + assumes "i\i'" shows "fmlookup (denvalue (updateEnvDup i t v e)) i' = fmlookup (denvalue e) i'" +proof (cases "fmlookup (denvalue e) i = None") + case True + then show ?thesis using assms by simp +next + case False + then obtain e' where "fmlookup (denvalue e) i = Some e'" by auto + then show ?thesis using assms by simp +qed + +lemma env_reorder_neq: + assumes "x\y" + shows "updateEnv x t1 v1 (updateEnv y t2 v2 e) = updateEnv y t2 v2 (updateEnv x t1 v1 e)" +proof - + have "address (updateEnv x t1 v1 (updateEnv y t2 v2 e)) = address (updateEnv y t2 v2 (updateEnv x t1 v1 e))" by simp + moreover from assms have "denvalue (updateEnv x t1 v1 (updateEnv y t2 v2 e)) = denvalue (updateEnv y t2 v2 (updateEnv x t1 v1 e))" using Finite_Map.fmupd_reorder_neq[of x y "(t1,v1)" "(t2,v2)"] by simp + ultimately show ?thesis by simp +qed + +lemma uEO_in: + assumes "i |\| fmdom (denvalue e)" + shows "updateEnvOption i t v e = None" + using assms by auto + +lemma uEO_n_In: + assumes "\ i |\| fmdom (denvalue e)" + shows "updateEnvOption i t v e = Some (updateEnv i t v e)" + using assms by auto + +fun astack :: "Identifier \ Type \ Stackvalue \ Stack * Environment \ Stack * Environment" + where "astack i t v (s, e) = (push v s, (updateEnv i t (Stackloc (ShowL\<^sub>n\<^sub>a\<^sub>t (toploc s))) e))" + +subsubsection \Examples\ +abbreviation "myenv::Environment \ eempty \denvalue := fmupd STR ''id1'' (Value TBool, Stackloc STR ''0'') fmempty\" +abbreviation "mystack::Stack \ \mapping = fmupd (STR ''0'') (KValue STR ''True'') fmempty, toploc = 1\" + +abbreviation "myenv2::Environment \ eempty \denvalue := fmupd STR ''id2'' (Value TBool, Stackloc STR ''1'') (fmupd STR ''id1'' (Value TBool, Stackloc STR ''0'') fmempty)\" +abbreviation "mystack2::Stack \ \mapping = fmupd (STR ''1'') (KValue STR ''False'') (fmupd (STR ''0'') (KValue STR ''True'') fmempty), toploc = 2\" + +lemma "astack (STR ''id1'') (Value TBool) (KValue (STR ''True'')) (emptyStore, eempty) = (mystack,myenv)" by solidity_symbex +lemma "astack (STR ''id2'') (Value TBool) (KValue (STR ''False'')) (mystack, myenv) = (mystack2,myenv2)" by solidity_symbex + +subsection \Declarations\ + +text \This function is used to declare a new variable: decl id tp val copy cd mem sto c m k e +\begin{description} + \item[id] is the name of the variable + \item[tp] is the type of the variable + \item[val] is an optional initialization parameter. If it is None, the types default value is taken. + \item[copy] is a flag to indicate whether memory should be copied (from mem parameter) or not (copying is required for example for external method calls). + \item[cd] is the original calldata which is used as a source + \item[mem] is the original memory which is used as a source + \item[sto] is the original storage which is used as a source + \item[c] is the new calldata which is updated + \item[m] is the new memory which is updated + \item[k] is the new calldata which is updated + \item[e] is the new environment which is updated +\end{description}\ +fun decl :: "Identifier \ Type \ (Stackvalue * Type) option \ bool \ CalldataT \ MemoryT \ (Address \ StorageT) + \ CalldataT \ MemoryT \ Stack \ Environment \ (CalldataT \ MemoryT \ Stack \ Environment) option" + where +(* Declaring stack variables *) + "decl i (Value t) None _ _ _ _ (c, m, k, e) = Some (c, m, (astack i (Value t) (KValue (ival t)) (k, e)))" +| "decl i (Value t) (Some (KValue v, Value t')) _ _ _ _ (c, m, k, e) = + Option.bind (convert t' t v) + (\v'. Some (c, m, astack i (Value t) (KValue v') (k, e)))" +| "decl _ (Value _) _ _ _ _ _ _ = None" + +(* Declaring calldata variables *) +| "decl i (Calldata (MTArray x t)) (Some (KCDptr p, _)) True cd _ _ (c, m, k, e) = + (let l = ShowL\<^sub>n\<^sub>a\<^sub>t (toploc c); + (_, c') = allocate c + in Option.bind (cpm2m p l x t cd c') + (\c''. Some (c'', m, astack i (Calldata (MTArray x t)) (KCDptr l) (k, e))))" +| "decl i (Calldata (MTArray x t)) (Some (KMemptr p, _)) True _ mem _ (c, m, k, e) = + (let l = ShowL\<^sub>n\<^sub>a\<^sub>t (toploc c); + (_, c') = allocate c + in Option.bind (cpm2m p l x t mem c') + (\c''. Some (c'', m, astack i (Calldata (MTArray x t)) (KCDptr l) (k, e))))" +| "decl i (Calldata _) _ _ _ _ _ _ = None" + +(* Declaring memory variables *) +| "decl i (Memory (MTArray x t)) None _ _ _ _ (c, m, k, e) = + (let m' = minit x t m + in Some (c, m', astack i (Memory (MTArray x t)) (KMemptr (ShowL\<^sub>n\<^sub>a\<^sub>t (toploc m))) (k, e)))" +| "decl i (Memory (MTArray x t)) (Some (KMemptr p, _)) True _ mem _ (c, m, k, e) = + Option.bind (cpm2m p (ShowL\<^sub>n\<^sub>a\<^sub>t (toploc m)) x t mem (snd (allocate m))) + (\m'. Some (c, m', astack i (Memory (MTArray x t)) (KMemptr (ShowL\<^sub>n\<^sub>a\<^sub>t (toploc m))) (k, e)))" +| "decl i (Memory (MTArray x t)) (Some (KMemptr p, _)) False _ _ _ (c, m, k, e) = + Some (c, m, astack i (Memory (MTArray x t)) (KMemptr p) (k, e))" +| "decl i (Memory (MTArray x t)) (Some (KCDptr p, _)) _ cd _ _ (c, m, k, e) = + Option.bind (cpm2m p (ShowL\<^sub>n\<^sub>a\<^sub>t (toploc m)) x t cd (snd (allocate m))) + (\m'. Some (c, m', astack i (Memory (MTArray x t)) (KMemptr (ShowL\<^sub>n\<^sub>a\<^sub>t (toploc m))) (k, e)))" +| "decl i (Memory (MTArray x t)) (Some (KStoptr p, Storage (STArray x' t'))) _ _ _ s (c, m, k, e) = + Option.bind (cps2m p (ShowL\<^sub>n\<^sub>a\<^sub>t (toploc m)) x' t' (s (address e)) (snd (allocate m))) + (\m''. Some (c, m'', astack i (Memory (MTArray x t)) (KMemptr (ShowL\<^sub>n\<^sub>a\<^sub>t (toploc m))) (k, e)))" +| "decl _ (Memory _) _ _ _ _ _ _ = None" + +(* Declaring storage variables *) +| "decl i (Storage (STArray x t)) (Some (KStoptr p, _)) _ _ _ _ (c, m, k, e) = + Some (c, m, astack i (Storage (STArray x t)) (KStoptr p) (k, e))" +| "decl i (Storage (STMap t t')) (Some (KStoptr p, _)) _ _ _ _ (c, m, k, e) = + Some (c, m, astack i (Storage (STMap t t')) (KStoptr p) (k, e))" +| "decl _ (Storage _) _ _ _ _ _ _ = None" (* Uninitialized storage arrays/maps not allowed *) + +lemma decl_env: + assumes "decl a1 a2 a3 cp cd mem sto (c, m, k, env) = Some (c', m', k', env')" + shows "address env = address env' \ sender env = sender env' \ svalue env = svalue env' \ (\x. x \ a1 \ fmlookup (denvalue env') x = fmlookup (denvalue env) x)" + using assms +proof (cases rule:decl.elims) + case (1 t uu uv uw ux c m k env) + then show ?thesis by simp +next + case (2 t v t' uy uz va vb c m k e) + show ?thesis + proof (cases "convert t' t v") + case None + with 2 show ?thesis by simp + next + case s: (Some a) + with 2 s show ?thesis by simp + qed +next + case (3 vd ve vb vf vg vh vi vj) + then show ?thesis by simp +next + case (4 vd ve vb vf vg vh vi vj) + then show ?thesis by simp +next + case (5 vd ve vb vf vg vh vi vj) + then show ?thesis by simp +next + case (6 vd va ve vf vg vh vi vj) + then show ?thesis by simp +next + case (7 vd va ve vf vg vh vi vj) + then show ?thesis by simp +next + case (8 vd va ve vf vg vh vi vj) + then show ?thesis by simp +next + case (9 x t p vk cd vl vm c m k env) + define l where "l = ShowL\<^sub>n\<^sub>a\<^sub>t (toploc c)" + obtain c' where c_def: "\x. (x, c') = allocate c" unfolding allocate_def by simp + show ?thesis + proof (cases "cpm2m p l x t cd c'") + case None + with 9 l_def c_def show ?thesis unfolding allocate_def by simp + next + case s2: (Some a) + with 9 l_def c_def show ?thesis unfolding allocate_def by simp + qed +next + case (10 x t p vn vo mem vp c m k env) + define l where "l = ShowL\<^sub>n\<^sub>a\<^sub>t (toploc c)" + obtain c' where c_def: "\x. (x, c') = allocate c" unfolding allocate_def by simp + show ?thesis + proof (cases "cpm2m p l x t mem c'") + case None + with 10 l_def c_def show ?thesis unfolding allocate_def by simp + next + case s2: (Some a) + with 10 l_def c_def show ?thesis unfolding allocate_def by simp + qed +next + case (11 v vr vs vt vu vv vw) + then show ?thesis by simp +next + case (12 vq vs vt vu vv vw) + then show ?thesis by simp +next + case (13 vq vc vb vs vt vu vv vw) + then show ?thesis by simp +next + case (14 v vc vb vs vt vu vv vw) + then show ?thesis by simp +next + case (15 vq vc vb vt vu vv vw) + then show ?thesis by simp +next + case (16 vq vc vb vs vt vu vv vw) + then show ?thesis by simp +next + case (17 vq vr vt vu vv vw) + then show ?thesis by simp +next + case (18 x t vx vy vz wa c m k env) + then show ?thesis by simp +next + case (19 x t p wb wc mem wd c m k env) + show ?thesis + proof (cases cp) + case True + define l where "l = ShowL\<^sub>n\<^sub>a\<^sub>t (toploc m)" + obtain m' where m'_def: "\x. (x, m') = allocate m" unfolding allocate_def by simp + then show ?thesis + proof (cases "cpm2m p l x t mem m'") + case None + with 19 l_def m'_def show ?thesis unfolding allocate_def by simp + next + case s2: (Some a) + with 19 l_def m'_def show ?thesis unfolding allocate_def by simp + qed + next + case False + with 19 show ?thesis by simp + qed +next + case (20 x t p we wf wg wh c m k env) + then show ?thesis by simp +next + case (21 x t p wi wj cd wk wl c m k env) + define l where "l = ShowL\<^sub>n\<^sub>a\<^sub>t (toploc m)" + obtain m' where m'_def: "\x. (x, m') = allocate m" unfolding allocate_def by simp + then show ?thesis + proof (cases "cpm2m p l x t cd m'") + case None + with 21 l_def m'_def show ?thesis unfolding allocate_def by simp + next + case s2: (Some a) + with 21 l_def m'_def show ?thesis unfolding allocate_def by simp + qed +next + case (22 x t p x' t' wm wn wo s c m k env) + define l where "l = ShowL\<^sub>n\<^sub>a\<^sub>t (toploc m)" + obtain m' where m'_def: "\x. (x, m') = allocate m" unfolding allocate_def by simp + show ?thesis + proof (cases "cps2m p l x' t' (s (address env)) m'") + case None + with 22 l_def m'_def show ?thesis unfolding allocate_def by simp + next + case s2: (Some a) + with 22 l_def m'_def show ?thesis unfolding allocate_def by simp + qed +next + case (23 v wr ws wt wu wv ww) + then show ?thesis by simp +next + case (24 va v ws wt wu wv ww) + then show ?thesis by simp +next + case (25 wq vc vb ws wt wu wv ww) + then show ?thesis by simp +next + case (26 v vc vb ws wt wu wv ww) + then show ?thesis by simp +next + case (27 v vc vb ws wt wu wv ww) + then show ?thesis by simp +next + case (28 wq vc v ws wt wu wv ww) + then show ?thesis by simp +next + case (29 wq vc v ws wt wu wv ww) + then show ?thesis by simp +next + case (30 wq vc v ws wt wu wv ww) + then show ?thesis by simp +next + case (31 wq vc va vd ws wt wu wv ww) + then show ?thesis by simp +next + case (32 wq vc va ws wt wu wv ww) + then show ?thesis by simp +next + case (33 va v wt wu wv ww) + then show ?thesis by simp +next + case (34 wq vc vb wt wu wv ww) + then show ?thesis by simp +next + case (35 v vc vb wt wu wv ww) + then show ?thesis by simp +next + case (36 v vc vb wt wu wv ww) + then show ?thesis by simp +next + case (37 wq vc v wt wu wv ww) + then show ?thesis by simp +next + case (38 wq vc v wt wu wv ww) + then show ?thesis by simp +next + case (39 wq vc v wt wu wv ww) + then show ?thesis by simp +next + case (40 wq vc va vd wt wu wv ww) + then show ?thesis by simp +next + case (41 wq vc va wt wu wv ww) + then show ?thesis by simp +next + case (42 x t p wx wy wz xa xb c m k env) + then show ?thesis by simp +next + case (43 t t' p xc xd xe xf xg c m k e) + then show ?thesis by simp +next + case (44 v va xk xl xm xn xo) + then show ?thesis by simp +next + case (45 v va ve vd xk xl xm xn xo) + then show ?thesis by simp +next + case (46 v va ve vd xk xl xm xn xo) + then show ?thesis by simp +next + case (47 v va ve vd xk xl xm xn xo) + then show ?thesis by simp +next + case (48 v xj xk xl xm xn xo) + then show ?thesis by simp +next + case (49 xi xk xl xm xn xo) + then show ?thesis by simp +next + case (50 xi vc vb xk xl xm xn xo) + then show ?thesis by simp +next + case (51 xi vc vb xk xl xm xn xo) + then show ?thesis by simp +next + case (52 xi vc vb xk xl xm xn xo) + then show ?thesis by simp +qed + +declare decl.simps[simp del] +end diff --git a/thys/Solidity/Expressions.thy b/thys/Solidity/Expressions.thy new file mode 100644 --- /dev/null +++ b/thys/Solidity/Expressions.thy @@ -0,0 +1,1376 @@ +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] + +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) + +datatype LType = LStackloc Location + | LMemloc Location + | LStoreloc Location + +locale expressions_with_gas = + fixes costs\<^sub>e :: "E \ Environment \ CalldataT \ State \ Gas" + and ep::Environment\<^sub>P + assumes call_not_zero[termination_simp]: "\e cd st i ix. 0 < (costs\<^sub>e (CALL i ix) e cd st)" + and ecall_not_zero[termination_simp]: "\e cd st a i ix. 0 < (costs\<^sub>e (ECALL a i ix) e cd st)" +begin +function (domintros) msel::"bool \ MTypes \ Location \ E list \ Environment \ CalldataT \ State \ (Location * MTypes, Ex, Gas) state_monad" + and ssel::"STypes \ Location \ E list \ Environment \ CalldataT \ State \ (Location * STypes, Ex, Gas) state_monad" + and expr::"E \ Environment \ CalldataT \ State \ (Stackvalue * Type, Ex, Gas) state_monad" + and load :: "bool \ (Identifier \ Type) list \ E list \ Environment \ CalldataT \ Stack \ MemoryT \ Environment \ CalldataT \ State \ (Environment \ CalldataT \ Stack \ MemoryT, Ex, Gas) state_monad" + and rexp::"L \ Environment \ CalldataT \ State \ (Stackvalue * Type, Ex, Gas) state_monad" +where + "msel _ _ _ [] _ _ _ g = throw Err g" +| "msel _ (MTValue _) _ _ _ _ _ g = throw Err g" +| "msel _ (MTArray al t) loc [x] env cd st g = + (do { + kv \ expr x env cd st; + (v, t') \ case kv of (KValue v, Value t') \ return (v, t') | _ \ throw Err; + assert Err (\_. less t' (TUInt 256) v (ShowL\<^sub>i\<^sub>n\<^sub>t al) = Some (ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True, TBool)); + return (hash loc v, t) + }) g" +(* + Note that it is indeed possible to modify the state while evaluating the expression + to determine the index of the array to look up. +*) +| "msel mm (MTArray al t) loc (x # y # ys) env cd st g = + (do { + kv \ expr x env cd st; + (v, t') \ case kv of (KValue v, Value t') \ return (v,t') | _ \ throw Err; + assert Err (\_. less t' (TUInt 256) v (ShowL\<^sub>i\<^sub>n\<^sub>t al) = Some (ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True, TBool)); + l \ case accessStore (hash loc v) (if mm then memory st else cd) of Some (MPointer l) \ return l | _ \ throw Err; + msel mm t l (y#ys) env cd st + }) g" +| "ssel tp loc Nil _ _ _ g = return (loc, tp) g" +| "ssel (STValue _) _ (_ # _) _ _ _ g = throw Err g" +| "ssel (STArray al t) loc (x # xs) env cd st g = + (do { + kv \ expr x env cd st; + (v, t') \ case kv of (KValue v, Value t') \ return (v, t') | _ \ throw Err; + assert Err (\_. less t' (TUInt 256) v (ShowL\<^sub>i\<^sub>n\<^sub>t al) = Some (ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True, TBool)); + ssel t (hash loc v) xs env cd st + }) g" +| "ssel (STMap _ t) loc (x # xs) env cd st g = + (do { + kv \ expr x env cd st; + v \ case kv of (KValue v, _) \ return v | _ \ throw Err; + ssel t (hash loc v) xs env cd st + }) g" +| "expr (E.INT b x) e cd st g = + (do { + assert Gas (\g. g > costs\<^sub>e (E.INT b x) e cd st); + modify (\g. g - costs\<^sub>e (E.INT b x) e cd st); + assert Err (\_. b \ vbits); + return (KValue (createSInt b x), Value (TSInt b)) + }) g" +| "expr (UINT b x) e cd st g = + (do { + assert Gas (\g. g > costs\<^sub>e (UINT b x) e cd st); + modify (\g. g - costs\<^sub>e (UINT b x) e cd st); + assert Err (\_. b \ vbits); + return (KValue (createUInt b x), Value (TUInt b)) + }) g" +| "expr (ADDRESS ad) e cd st g = + (do { + assert Gas (\g. g > costs\<^sub>e (ADDRESS ad) e cd st); + modify (\g. g - costs\<^sub>e (ADDRESS ad) e cd st); + return (KValue ad, Value TAddr) + }) g" +| "expr (BALANCE ad) e cd st g = + (do { + assert Gas (\g. g > costs\<^sub>e (BALANCE ad) e cd st); + modify (\g. g - costs\<^sub>e (BALANCE ad) e cd st); + kv \ expr ad e cd st; + adv \ case kv of (KValue adv, Value TAddr) \ return adv | _ \ throw Err; + return (KValue (bal ((accounts st) adv)), Value (TUInt 256)) + }) g" +| "expr THIS e cd st g = + (do { + assert Gas (\g. g > costs\<^sub>e THIS e cd st); + modify (\g. g - costs\<^sub>e THIS e cd st); + return (KValue (address e), Value TAddr) + }) g" +| "expr SENDER e cd st g = + (do { + assert Gas (\g. g > costs\<^sub>e SENDER e cd st); + modify (\g. g - costs\<^sub>e SENDER e cd st); + return (KValue (sender e), Value TAddr) + }) g" +| "expr VALUE e cd st g = + (do { + assert Gas (\g. g > costs\<^sub>e VALUE e cd st); + modify (\g. g - costs\<^sub>e VALUE e cd st); + return (KValue (svalue e), Value (TUInt 256)) + }) g" +| "expr TRUE e cd st g = + (do { + assert Gas (\g. g > costs\<^sub>e TRUE e cd st); + modify (\g. g - costs\<^sub>e TRUE e cd st); + return (KValue (ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True), Value TBool) + }) g" +| "expr FALSE e cd st g = + (do { + assert Gas (\g. g > costs\<^sub>e FALSE e cd st); + modify (\g. g - costs\<^sub>e FALSE e cd st); + return (KValue (ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l False), Value TBool) + }) g" +| "expr (NOT x) e cd st g = + (do { + assert Gas (\g. g > costs\<^sub>e (NOT x) e cd st); + modify (\g. g - costs\<^sub>e (NOT x) e cd st); + kv \ expr x e cd st; + v \ case kv of (KValue v, Value TBool) \ return v | _ \ throw Err; + (if v = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True then expr FALSE e cd st + else if v = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l False then expr TRUE e cd st + else throw Err) + }) g" +| "expr (PLUS e1 e2) e cd st g = + (do { + assert Gas (\g. g > costs\<^sub>e (PLUS e1 e2) e cd st); + modify (\g. g - costs\<^sub>e (PLUS e1 e2) e cd st); + lift expr add e1 e2 e cd st + }) g" +| "expr (MINUS e1 e2) e cd st g = + (do { + assert Gas (\g. g > costs\<^sub>e (MINUS e1 e2) e cd st); + modify (\g. g - costs\<^sub>e (MINUS e1 e2) e cd st); + lift expr sub e1 e2 e cd st + }) g" +| "expr (LESS e1 e2) e cd st g = + (do { + assert Gas (\g. g > costs\<^sub>e (LESS e1 e2) e cd st); + modify (\g. g - costs\<^sub>e (LESS e1 e2) e cd st); + lift expr less e1 e2 e cd st + }) g" +| "expr (EQUAL e1 e2) e cd st g = + (do { + assert Gas (\g. g > costs\<^sub>e (EQUAL e1 e2) e cd st); + modify (\g. g - costs\<^sub>e (EQUAL e1 e2) e cd st); + lift expr equal e1 e2 e cd st + }) g" +| "expr (AND e1 e2) e cd st g = + (do { + assert Gas (\g. g > costs\<^sub>e (AND e1 e2) e cd st); + modify (\g. g - costs\<^sub>e (AND e1 e2) e cd st); + lift expr vtand e1 e2 e cd st + }) g" +| "expr (OR e1 e2) e cd st g = + (do { + assert Gas (\g. g > costs\<^sub>e (OR e1 e2) e cd st); + modify (\g. g - costs\<^sub>e (OR e1 e2) e cd st); + lift expr vtor e1 e2 e cd st + }) g" +| "expr (LVAL i) e cd st g = + (do { + assert Gas (\g. g > costs\<^sub>e (LVAL i) e cd st); + modify (\g. g - costs\<^sub>e (LVAL i) e cd st); + rexp i e cd st + }) g" +(* Notes about method calls: + - Internal method calls use a fresh environment and stack but keep the memory [1] + - External method calls use a fresh environment and stack and memory [2] + [1]: https://docs.soliditylang.org/en/v0.8.5/control-structures.html#internal-function-calls + [2]: https://docs.soliditylang.org/en/v0.8.5/control-structures.html#external-function-calls +*) +| "expr (CALL i xe) e cd st g = + (do { + assert Gas (\g. g > costs\<^sub>e (CALL i xe) e cd st); + modify (\g. g - costs\<^sub>e (CALL i xe) e cd st); + (ct, _) \ option Err (\_. ep $$ (contract e)); + (fp, x) \ case ct $$ i of Some (Function (fp, False, x)) \ return (fp, x) | _ \ throw Err; + let e' = ffold_init ct (emptyEnv (address e) (contract e) (sender e) (svalue e)) (fmdom ct); + (e\<^sub>l, cd\<^sub>l, k\<^sub>l, m\<^sub>l) \ load False fp xe e' emptyStore emptyStore (memory st) e cd st; + expr x e\<^sub>l cd\<^sub>l (st\stack:=k\<^sub>l, memory:=m\<^sub>l\) + }) g" +(*It is not allowed to transfer money to external function calls*) +| "expr (ECALL ad i xe) e cd st g = + (do { + assert Gas (\g. g > costs\<^sub>e (ECALL ad i xe) e cd st); + modify (\g. g - costs\<^sub>e (ECALL ad i xe) e cd st); + kad \ expr ad e cd st; + adv \ case kad of (KValue adv, Value TAddr) \ return adv | _ \ throw Err; + assert Err (\_. adv \ address e); + c \ case type (accounts st adv) of Some (Contract c) \ return c | _ \ throw Err; + (ct, _) \ option Err (\_. ep $$ c); + (fp, x) \ case ct $$ i of Some (Function (fp, True, x)) \ return (fp, x) | _ \ throw Err; + let e' = ffold_init ct (emptyEnv adv c (address e) (ShowL\<^sub>n\<^sub>a\<^sub>t 0)) (fmdom ct); + (e\<^sub>l, cd\<^sub>l, k\<^sub>l, m\<^sub>l) \ load True fp xe e' emptyStore emptyStore emptyStore e cd st; + expr x e\<^sub>l cd\<^sub>l (st\stack:=k\<^sub>l, memory:=m\<^sub>l\) + }) g" +| "load cp ((i\<^sub>p, t\<^sub>p)#pl) (ex#el) e\<^sub>v' cd' sck' mem' e\<^sub>v cd st g = + (do { + (v, t) \ expr ex e\<^sub>v cd st; + (c, m, k, e) \ case decl i\<^sub>p t\<^sub>p (Some (v,t)) cp cd (memory st) (storage st) (cd', mem', sck', e\<^sub>v') of Some (c, m, k, e) \ return (c, m, k, e) | None \ throw Err; + load cp pl el e c k m e\<^sub>v cd st + }) g" +| "load _ [] (_#_) _ _ _ _ _ _ _ g = throw Err g" +| "load _ (_#_) [] _ _ _ _ _ _ _ g = throw Err g" +| "load _ [] [] e\<^sub>v' cd' sck' mem' e\<^sub>v cd st g = return (e\<^sub>v', cd', sck', mem') g" + +| "rexp (Id i) e cd st g = + (case fmlookup (denvalue e) i of + Some (tp, Stackloc l) \ + (case accessStore l (stack st) of + Some (KValue v) \ return (KValue v, tp) + | Some (KCDptr p) \ return (KCDptr p, tp) + | Some (KMemptr p) \ return (KMemptr p, tp) + | Some (KStoptr p) \ return (KStoptr p, tp) + | _ \ throw Err) + | Some (Storage (STValue t), Storeloc l) \ return (KValue (accessStorage t l (storage st (address e))), Value t) + | Some (Storage (STArray x t), Storeloc l) \ return (KStoptr l, Storage (STArray x t)) + | _ \ throw Err) g" +| "rexp (Ref i r) e cd st g = + (case fmlookup (denvalue e) i of + Some (tp, (Stackloc l)) \ + (case accessStore l (stack st) of + Some (KCDptr l') \ + do { + t \ case tp of Calldata t \ return t | _ \ throw Err; + (l'', t') \ msel False t l' r e cd st; + (case t' of + MTValue t'' \ + do { + v \ case accessStore l'' cd of Some (MValue v) \ return v | _ \ throw Err; + return (KValue v, Value t'') + } + | MTArray x t'' \ + do { + p \ case accessStore l'' cd of Some (MPointer p) \ return p | _ \ throw Err; + return (KCDptr p, Calldata (MTArray x t'')) + } + ) + } + | Some (KMemptr l') \ + do { + t \ case tp of Memory t \ return t | _ \ throw Err; + (l'', t') \ msel True t l' r e cd st; + (case t' of + MTValue t'' \ + do { + v \ case accessStore l'' (memory st) of Some (MValue v) \ return v | _ \ throw Err; + return (KValue v, Value t'') + } + | MTArray x t'' \ + do { + p \ case accessStore l'' (memory st) of Some (MPointer p) \ return p | _ \ throw Err; + return (KMemptr p, Memory (MTArray x t'')) + } + ) + } + | Some (KStoptr l') \ + do { + t \ case tp of Storage t \ return t | _ \ throw Err; + (l'', t') \ ssel t l' r e cd st; + (case t' of + STValue t'' \ return (KValue (accessStorage t'' l'' (storage st (address e))), Value t'') + | STArray _ _ \ return (KStoptr l'', Storage t') + | STMap _ _ \ return (KStoptr l'', Storage t')) + } + | _ \ throw Err) + | Some (tp, (Storeloc l)) \ + do { + t \ case tp of Storage t \ return t | _ \ throw Err; + (l', t') \ ssel t l r e cd st; + (case t' of + STValue t'' \ return (KValue (accessStorage t'' l' (storage st (address e))), Value t'') + | STArray _ _ \ return (KStoptr l', Storage t') + | STMap _ _ \ return (KStoptr l', Storage t')) + } + | None \ throw Err) g" +| "expr CONTRACTS e cd st g = + (do { + assert Gas (\g. g > costs\<^sub>e CONTRACTS e cd st); + modify (\g. g - costs\<^sub>e CONTRACTS e cd st); + prev \ case contracts (accounts st (address e)) of 0 \ throw Err | Suc n \ return n; + return (KValue (hash (address e) (ShowL\<^sub>n\<^sub>a\<^sub>t prev)), Value TAddr) + }) g" + by pat_completeness auto + +subsection \Termination\ + +text \To prove termination we first need to show that expressions do not increase gas\ + +lemma lift_gas: + assumes "lift expr f e1 e2 e cd st g = Normal (v, g')" + and "\v g'. expr e1 e cd st g = Normal (v, g') \ g' \ g" + and "\v g' v' t' g''. expr e1 e cd st g = Normal (v, g') + \ expr e2 e cd st g' = Normal (v', g'') + \ g'' \ g'" + shows "g' \ g" +proof (cases "expr e1 e cd st g") + case (n a g0') + then show ?thesis + proof (cases a) + case (Pair b c) + then show ?thesis + proof (cases b) + case (KValue v1) + then show ?thesis + proof (cases c) + case (Value t1) + then show ?thesis + proof (cases "expr e2 e cd st g0'") + case r2: (n a' g0'') + then show ?thesis + proof (cases a') + case p2: (Pair b c) + then show ?thesis + proof (cases b) + case v2: (KValue v2) + then show ?thesis + proof (cases c) + case t2: (Value t2) + then show ?thesis + proof (cases "f t1 t2 v1 v2") + case None + with assms n Pair KValue Value r2 p2 v2 t2 show ?thesis by simp + next + case (Some a'') + then show ?thesis + proof (cases a'') + case p3: (Pair v t) + with assms n Pair KValue Value r2 p2 v2 t2 Some have "g0'\ g" by simp + moreover from assms n Pair KValue Value r2 p2 v2 t2 Some have "g0''\g0'" by simp + moreover from assms n Pair KValue Value r2 p2 v2 t2 Some have "g'=g0''" by (simp split:prod.split_asm) + ultimately show ?thesis by arith + qed + qed + next + case (Calldata x2) + with assms n Pair KValue Value r2 p2 v2 show ?thesis by simp + next + case (Memory x3) + with assms n Pair KValue Value r2 p2 v2 show ?thesis by simp + next + case (Storage x4) + with assms n Pair KValue Value r2 p2 v2 show ?thesis by simp + qed + next + case (KCDptr x2) + with assms n Pair KValue Value r2 p2 show ?thesis by simp + next + case (KMemptr x3) + with assms n Pair KValue Value r2 p2 show ?thesis by simp + next + case (KStoptr x4) + with assms n Pair KValue Value r2 p2 show ?thesis by simp + qed + qed + next + case (e x) + with assms n Pair KValue Value show ?thesis by simp + qed + next + case (Calldata x2) + with assms n Pair KValue show ?thesis by simp + next + case (Memory x3) + with assms n Pair KValue show ?thesis by simp + next + case (Storage x4) + with assms n Pair KValue show ?thesis by simp + qed + next + case (KCDptr x2) + with assms n Pair show ?thesis by simp + next + case (KMemptr x3) + with assms n Pair show ?thesis by simp + next + case (KStoptr x4) + with assms n Pair show ?thesis by simp + qed + qed +next + case (e x) + with assms show ?thesis by simp +qed + +lemma msel_ssel_expr_load_rexp_dom_gas[rule_format]: + "msel_ssel_expr_load_rexp_dom (Inl (Inl (c1, t1, l1, xe1, ev1, cd1, st1, g1))) + \ (\v1' g1'. msel c1 t1 l1 xe1 ev1 cd1 st1 g1 = Normal (v1', g1') \ g1' \ g1)" + "msel_ssel_expr_load_rexp_dom (Inl (Inr (t2, l2, xe2, ev2, cd2, st2, g2))) + \ (\v2' g2'. ssel t2 l2 xe2 ev2 cd2 st2 g2 = Normal (v2', g2') \ g2' \ g2)" + "msel_ssel_expr_load_rexp_dom (Inr (Inl (e4, ev4, cd4, st4, g4))) + \ (\v4' g4'. expr e4 ev4 cd4 st4 g4 = Normal (v4', g4') \ g4' \ g4)" + "msel_ssel_expr_load_rexp_dom (Inr (Inr (Inl (lcp, lis, lxs, lev0, lcd0, lk, lm, lev, lcd, lst, lg)))) + \ (\ev cd k m g'. load lcp lis lxs lev0 lcd0 lk lm lev lcd lst lg = Normal ((ev, cd, k, m), g') \ g' \ lg \ address ev = address lev0 \ sender ev = sender lev0 \ svalue ev = svalue lev0)" + "msel_ssel_expr_load_rexp_dom (Inr (Inr (Inr (l3, ev3, cd3, st3, g3)))) + \ (\v3' g3'. rexp l3 ev3 cd3 st3 g3 = Normal (v3', g3') \ g3' \ g3)" +proof (induct rule: msel_ssel_expr_load_rexp.pinduct +[where ?P1.0="\c1 t1 l1 xe1 ev1 cd1 st1 g1. (\l1' g1'. msel c1 t1 l1 xe1 ev1 cd1 st1 g1 = Normal (l1', g1') \ g1' \ g1)" + and ?P2.0="\t2 l2 xe2 ev2 cd2 st2 g2. (\v2' g2'. ssel t2 l2 xe2 ev2 cd2 st2 g2 = Normal (v2', g2') \ g2' \ g2)" + and ?P3.0="\e4 ev4 cd4 st4 g4. (\v4' g4'. expr e4 ev4 cd4 st4 g4 = Normal (v4', g4') \ g4' \ g4)" + and ?P4.0="\lcp lis lxs lev0 lcd0 lk lm lev lcd lst lg. (\ev cd k m g'. load lcp lis lxs lev0 lcd0 lk lm lev lcd lst lg = Normal ((ev, cd, k, m), g') \ g' \ lg \ address ev = address lev0 \ sender ev = sender lev0 \ svalue ev = svalue lev0)" + and ?P5.0="\l3 ev3 cd3 st3 g3. (\v3' g3'. rexp l3 ev3 cd3 st3 g3 = Normal (v3', g3') \ g3' \ g3)" +]) + case 1 + then show ?case using msel.psimps(1) by auto +next + case 2 + then show ?case using msel.psimps(2) by auto +next + case 3 + then show ?case using msel.psimps(3) by (auto split: if_split_asm Type.split_asm Stackvalue.split_asm prod.split_asm StateMonad.result.split_asm) +next + case (4 mm al t loc x y ys env cd st g) + show ?case + proof (rule allI[THEN allI, OF impI]) + fix v1' g1' assume a1: "msel mm (MTArray al t) loc (x # y # ys) env cd st g = Normal (v1', g1')" + show "g1' \ g" + proof (cases v1') + case (Pair l1' t1') + then show ?thesis + proof (cases "expr x env cd st g") + case (n a g') + then show ?thesis + proof (cases a) + case p2: (Pair b c) + then show ?thesis + proof (cases b) + case (KValue v) + then show ?thesis + proof (cases c) + case (Value t') + then show ?thesis + proof (cases) + assume l: "less t' (TUInt 256) v (ShowL\<^sub>i\<^sub>n\<^sub>t al) = Some (ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True, TBool)" + then show ?thesis + proof (cases "accessStore (hash loc v) (if mm then memory st else cd)") + case None + with 4 a1 n p2 KValue Value l show ?thesis using msel.psimps(4) by simp + next + case (Some a) + then show ?thesis + proof (cases a) + case (MValue _) + with 4 a1 n p2 KValue Value Some l show ?thesis using msel.psimps(4) by simp + next + case (MPointer l) + with n p2 KValue Value l Some + have "msel mm (MTArray al t) loc (x # y # ys) env cd st g = msel mm t l (y # ys) env cd st g'" + using msel.psimps(4) 4(1) by simp + moreover from n have "g' \ g" using 4(2) by simp + moreover from a1 MPointer n Pair p2 KValue Value l Some + have "g1' \ g'" using msel.psimps(4) 4(3) 4(1) by simp + ultimately show ?thesis by simp + qed + qed + next + assume "\ less t' (TUInt 256) v (ShowL\<^sub>i\<^sub>n\<^sub>t al) = Some (ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True, TBool)" + with 4 a1 n p2 KValue Value show ?thesis using msel.psimps(4) by simp + qed + next + case (Calldata _) + with 4 a1 n p2 KValue show ?thesis using msel.psimps(4) by simp + next + case (Memory _) + with 4 a1 n p2 KValue show ?thesis using msel.psimps(4) by simp + next + case (Storage _) + with 4 a1 n p2 KValue show ?thesis using msel.psimps(4) by simp + qed + next + case (KCDptr _) + with 4 a1 n p2 show ?thesis using msel.psimps(4) by simp + next + case (KMemptr _) + with 4 a1 n p2 show ?thesis using msel.psimps(4) by simp + next + case (KStoptr _) + with 4 a1 n p2 show ?thesis using msel.psimps(4) by simp + qed + qed + next + case (e _) + with 4 a1 show ?thesis using msel.psimps(4) by simp + qed + qed + qed +next + case 5 + then show ?case using ssel.psimps(1) by auto +next + case 6 + then show ?case using ssel.psimps(2) by auto +next + case (7 al t loc x xs env cd st g) + show ?case + proof (rule allI[THEN allI, OF impI]) + fix v2' g2' assume a1: "ssel (STArray al t) loc (x # xs) env cd st g = Normal (v2', g2')" + show "g2' \ g" + proof (cases v2') + case (Pair l2' t2') + then show ?thesis + proof (cases "expr x env cd st g") + case (n a g'') + then show ?thesis + proof (cases a) + case p2: (Pair b c) + then show ?thesis + proof (cases b) + case (KValue v) + then show ?thesis + proof (cases c) + case (Value t') + then show ?thesis + proof (cases) + assume l: "less t' (TUInt 256) v (ShowL\<^sub>i\<^sub>n\<^sub>t al) = Some (ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True, TBool)" + with n p2 KValue Value l + have "ssel (STArray al t) loc (x # xs) env cd st g = ssel t (hash loc v) xs env cd st g''" + using ssel.psimps(3) 7(1) by simp + moreover from n have "g'' \ g" using 7(2) by simp + moreover from a1 n Pair p2 KValue Value l + have "g2' \ g''" using ssel.psimps(3) 7(3) 7(1) by simp + ultimately show ?thesis by simp + next + assume "\ less t' (TUInt 256) v (ShowL\<^sub>i\<^sub>n\<^sub>t al) = Some (ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True, TBool)" + with 7 a1 n p2 KValue Value show ?thesis using ssel.psimps(3) by simp + qed + next + case (Calldata _) + with 7 a1 n p2 KValue show ?thesis using ssel.psimps(3) by simp + next + case (Memory _) + with 7 a1 n p2 KValue show ?thesis using ssel.psimps(3) by simp + next + case (Storage _) + with 7 a1 n p2 KValue show ?thesis using ssel.psimps(3) by simp + qed + next + case (KCDptr _) + with 7 a1 n p2 show ?thesis using ssel.psimps(3) by simp + next + case (KMemptr _) + with 7 a1 n p2 show ?thesis using ssel.psimps(3) by simp + next + case (KStoptr x4) + with 7 a1 n p2 show ?thesis using ssel.psimps(3) by simp + qed + qed + next + case (e _) + with 7 a1 show ?thesis using ssel.psimps(3) by simp + qed + qed + qed +next + case (8 vv t loc x xs env cd st g) + show ?case + proof (rule allI[THEN allI, OF impI]) + fix v2' g2' assume a1: "ssel (STMap vv t) loc (x # xs) env cd st g = Normal (v2', g2')" + show "g2' \ g" + proof (cases v2') + case (Pair l2' t2') + then show ?thesis + proof (cases "expr x env cd st g") + case (n a g') + then show ?thesis + proof (cases a) + case p2: (Pair b c) + then show ?thesis + proof (cases b) + case (KValue v) + with 8 n p2 have "ssel (STMap vv t) loc (x # xs) env cd st g = ssel t (hash loc v) xs env cd st g'" using ssel.psimps(4) by simp + moreover from n have "g' \ g" using 8(2) by simp + moreover from a1 n Pair p2 KValue + have "g2' \ g'" using ssel.psimps(4) 8(3) 8(1) by simp + ultimately show ?thesis by simp + next + case (KCDptr _) + with 8 a1 n p2 show ?thesis using ssel.psimps(4) by simp + next + case (KMemptr _) + with 8 a1 n p2 show ?thesis using ssel.psimps(4) by simp + next + case (KStoptr _) + with 8 a1 n p2 show ?thesis using ssel.psimps(4) by simp + qed + qed + next + case (e _) + with 8 a1 show ?thesis using ssel.psimps(4) by simp + qed + qed + qed +next + case 9 + then show ?case using expr.psimps(1) by (simp split:if_split_asm) +next + case 10 + then show ?case using expr.psimps(2) by (simp split:if_split_asm) +next + case 11 + then show ?case using expr.psimps(3) by simp +next + case (12 ad e cd st g) + define gc where "gc = costs\<^sub>e (BALANCE ad) e cd st" + show ?case + proof (rule allI[THEN allI, OF impI]) + fix g4 xa + assume *: "expr (BALANCE ad) e cd st g = Normal (xa, g4)" + show "g4 \ g" + proof (cases) + assume "g \ gc" + with 12 gc_def * show ?thesis using expr.psimps(4) by simp + next + assume gcost: "\ g \ gc" + then show ?thesis + proof (cases "expr ad e cd st (g - gc)") + case (n a s) + show ?thesis + proof (cases a) + case (Pair b c) + then show ?thesis + proof (cases b) + case (KValue x1) + then show ?thesis + proof (cases c) + case (Value x1) + then show ?thesis + proof (cases x1) + case (TSInt _) + with 12 gc_def * gcost n Pair KValue Value show ?thesis using expr.psimps(4)[of ad e cd st] by simp + next + case (TUInt _) + with 12 gc_def * gcost n Pair KValue Value show ?thesis using expr.psimps(4)[of ad e cd st] by simp + next + case TBool + with 12 gc_def * gcost n Pair KValue Value show ?thesis using expr.psimps(4)[of ad e cd st] by simp + next + case TAddr + with 12(2)[where ?s'a="g-costs\<^sub>e (BALANCE ad) e cd st"] gc_def * gcost n Pair KValue Value show "g4 \ g" using expr.psimps(4)[OF 12(1)] by simp + qed + next + case (Calldata _) + with 12 gc_def * gcost n Pair KValue show ?thesis using expr.psimps(4)[of ad e cd st] by simp + next + case (Memory _) + with 12 gc_def * gcost n Pair KValue show ?thesis using expr.psimps(4)[of ad e cd st] by simp + next + case (Storage _) + with 12 gc_def * gcost n Pair KValue show ?thesis using expr.psimps(4)[of ad e cd st] by simp + qed + next + case (KCDptr _) + with 12 gc_def * gcost n Pair show ?thesis using expr.psimps(4)[of ad e cd st] by simp + next + case (KMemptr _) + with 12 gc_def * gcost n Pair show ?thesis using expr.psimps(4)[of ad e cd st] by simp + next + case (KStoptr _) + with 12 gc_def * gcost n Pair show ?thesis using expr.psimps(4)[of ad e cd st] by simp + qed + qed + next + case (e _) + with 12 gc_def * gcost show ?thesis using expr.psimps(4)[of ad e cd st] by simp + qed + qed + qed +next + case 13 + then show ?case using expr.psimps(5) by simp +next + case 14 + then show ?case using expr.psimps(6) by simp +next + case 15 + then show ?case using expr.psimps(7) by simp +next + case 16 + then show ?case using expr.psimps(8) by simp +next + case 17 + then show ?case using expr.psimps(9) by simp +next + case (18 x e cd st g) + define gc where "gc = costs\<^sub>e (NOT x) e cd st" + show ?case + proof (rule allI[THEN allI, OF impI]) + fix v4 g4' assume a1: "expr (NOT x) e cd st g = Normal (v4, g4')" + show "g4' \ g" + proof (cases v4) + case (Pair l4 t4) + show ?thesis + proof (cases) + assume "g \ gc" + with gc_def a1 show ?thesis using expr.psimps(10)[OF 18(1)] by simp + next + assume gcost: "\ g \ gc" + then show ?thesis + proof (cases "expr x e cd st (g - gc)") + case (n a g') + then show ?thesis + proof (cases a) + case p2: (Pair b c) + then show ?thesis + proof (cases b) + case (KValue v) + then show ?thesis + proof (cases c) + case (Value t) + then show ?thesis + proof (cases t) + case (TSInt x1) + with a1 gc_def gcost n p2 KValue Value show ?thesis using expr.psimps(10)[OF 18(1)] by simp + next + case (TUInt x2) + with a1 gc_def gcost n p2 KValue Value show ?thesis using expr.psimps(10)[OF 18(1)] by simp + next + case TBool + then show ?thesis + proof (cases) + assume v_def: "v = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True" + with 18(1) gc_def gcost n p2 KValue Value TBool have "expr (NOT x) e cd st g = expr FALSE e cd st g' " using expr.psimps(10)[OF 18(1)] by simp + moreover from 18(2) gc_def gcost n p2 have "g' \ g-gc" by simp + moreover from 18(3)[OF _ _ n] a1 gc_def gcost have "g4' \ g'" using expr.psimps(10)[OF 18(1)] n Pair p2 KValue Value TBool v_def by simp + ultimately show ?thesis by arith + next + assume v_def: "\ v = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True" + then show ?thesis + proof (cases) + assume v_def2: "v = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l False" + with 18(1) gc_def gcost n p2 KValue Value v_def TBool have "expr (NOT x) e cd st g = expr TRUE e cd st g'" using expr.psimps(10) by simp + moreover from 18(2)[where ?s'a="g-costs\<^sub>e (NOT x) e cd st"] gc_def gcost n p2 have "g' \ g" by simp + moreover from 18(4)[OF _ _ n] a1 gc_def gcost have "g4' \ g'" using expr.psimps(10)[OF 18(1)] n Pair p2 KValue Value TBool v_def v_def2 by simp + ultimately show ?thesis by arith + next + assume "\ v = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l False" + with 18(1) a1 gc_def gcost n p2 KValue Value v_def TBool show ?thesis using expr.psimps(10) by simp + qed + qed + next + case TAddr + with 18(1) a1 gc_def gcost n p2 KValue Value show ?thesis using expr.psimps(10) by simp + qed + next + case (Calldata _) + with 18(1) a1 gc_def gcost n p2 KValue show ?thesis using expr.psimps(10) by simp + next + case (Memory _) + with 18(1) a1 gc_def gcost n p2 KValue show ?thesis using expr.psimps(10) by simp + next + case (Storage _) + with 18(1) a1 gc_def gcost n p2 KValue show ?thesis using expr.psimps(10) by simp + qed + next + case (KCDptr _) + with 18(1) a1 gc_def gcost n p2 show ?thesis using expr.psimps(10) by simp + next + case (KMemptr _) + with 18(1) a1 gc_def gcost n p2 show ?thesis using expr.psimps(10) by simp + next + case (KStoptr _) + with 18(1) a1 gc_def gcost n p2 show ?thesis using expr.psimps(10) by simp + qed + qed + next + case (e _) + with 18(1) a1 gc_def gcost show ?thesis using expr.psimps(10) by simp + qed + qed + qed + qed +next + case (19 e1 e2 e cd st g) + define gc where "gc = costs\<^sub>e (PLUS e1 e2) e cd st" + show ?case + proof (rule allI[THEN allI, OF impI]) + fix g4 xa assume e_def: "expr (PLUS e1 e2) e cd st g = Normal (xa, g4)" + then show "g4 \ g" + proof (cases) + assume "g \ gc" + with 19(1) e_def show ?thesis using expr.psimps(11) gc_def by simp + next + assume "\ g \ gc" + then have *: "assert Gas ((<) (costs\<^sub>e (PLUS e1 e2) e cd st)) g = Normal ((), g)" using gc_def by simp + with 19(1) e_def gc_def have lift:"lift expr add e1 e2 e cd st (g - gc) = Normal (xa, g4)" using expr.psimps(11)[OF 19(1)] by simp + moreover have **: "modify (\g. g - costs\<^sub>e (PLUS e1 e2) e cd st) g = Normal ((), g - costs\<^sub>e (PLUS e1 e2) e cd st)" by simp + with 19(2)[OF * **] have "\g4' v4 t4. expr e1 e cd st (g-gc) = Normal ((v4, t4), g4') \ g4' \ g - gc" unfolding gc_def by simp + moreover obtain v g' where ***: "expr e1 e cd st (g - costs\<^sub>e (PLUS e1 e2) e cd st) = Normal (v, g')" using expr.psimps(11)[OF 19(1)] e_def by (simp split:if_split_asm result.split_asm prod.split_asm) + with 19(3)[OF * ** ***] have "\v t g' v' t' g''. expr e1 e cd st (g-gc) = Normal ((KValue v, Value t), g') \ expr e2 e cd st g' = Normal ((v', t'), g'') \ g'' \ g'" unfolding gc_def by simp + ultimately show "g4 \ g" using lift_gas[OF lift] `\ g \ gc` by (simp split:result.split_asm Stackvalue.split_asm Type.split_asm prod.split_asm) + qed + qed +next + case (20 e1 e2 e cd st g) + define gc where "gc = costs\<^sub>e (MINUS e1 e2) e cd st" + show ?case + proof (rule allI[THEN allI, OF impI]) + fix g4 xa assume e_def: "expr (MINUS e1 e2) e cd st g = Normal (xa, g4)" + then show "g4 \ g" + proof (cases) + assume "g \ gc" + with 20(1) e_def show ?thesis using expr.psimps(12) gc_def by simp + next + assume "\ g \ gc" + then have *: "assert Gas ((<) (costs\<^sub>e (MINUS e1 e2) e cd st)) g = Normal ((), g)" using gc_def by simp + with 20(1) e_def gc_def have lift:"lift expr sub e1 e2 e cd st (g - gc) = Normal (xa, g4)" using expr.psimps(12)[OF 20(1)] by simp + moreover have **: "modify (\g. g - costs\<^sub>e (MINUS e1 e2) e cd st) g = Normal ((), g - costs\<^sub>e (MINUS e1 e2) e cd st)" by simp + with 20(2)[OF * **] have "\g4' v4 t4. expr e1 e cd st (g-gc) = Normal ((v4, t4), g4') \ g4' \ g - gc" unfolding gc_def by simp + moreover obtain v g' where ***: "expr e1 e cd st (g - costs\<^sub>e (MINUS e1 e2) e cd st) = Normal (v, g')" using expr.psimps(12)[OF 20(1)] e_def by (simp split:if_split_asm result.split_asm prod.split_asm) + with 20(3)[OF * ** ***] have "\v t g' v' t' g''. expr e1 e cd st (g-gc) = Normal ((KValue v, Value t), g') \ expr e2 e cd st g' = Normal ((v', t'), g'') \ g'' \ g'" unfolding gc_def by simp + ultimately show "g4 \ g" using lift_gas[OF lift] `\ g \ gc` by (simp split:result.split_asm Stackvalue.split_asm Type.split_asm prod.split_asm) + qed + qed +next + case (21 e1 e2 e cd st g) + define gc where "gc = costs\<^sub>e (LESS e1 e2) e cd st" + show ?case + proof (rule allI[THEN allI, OF impI]) + fix g4 xa assume e_def: "expr (LESS e1 e2) e cd st g = Normal (xa, g4)" + then show "g4 \ g" + proof (cases) + assume "g \ gc" + with 21(1) e_def show ?thesis using expr.psimps(13) gc_def by simp + next + assume "\ g \ gc" + then have *: "assert Gas ((<) (costs\<^sub>e (LESS e1 e2) e cd st)) g = Normal ((), g)" using gc_def by simp + with 21(1) e_def gc_def have lift:"lift expr less e1 e2 e cd st (g - gc) = Normal (xa, g4)" using expr.psimps(13)[OF 21(1)] by simp + moreover have **: "modify (\g. g - costs\<^sub>e (LESS e1 e2) e cd st) g = Normal ((), g - costs\<^sub>e (LESS e1 e2) e cd st)" by simp + with 21(2)[OF * **] have "\g4' v4 t4. expr e1 e cd st (g-gc) = Normal ((v4, t4), g4') \ g4' \ g - gc" unfolding gc_def by simp + moreover obtain v g' where ***: "expr e1 e cd st (g - costs\<^sub>e (LESS e1 e2) e cd st) = Normal (v, g')" using expr.psimps(13)[OF 21(1)] e_def by (simp split:if_split_asm result.split_asm prod.split_asm) + with 21(3)[OF * ** ***] have "\v t g' v' t' g''. expr e1 e cd st (g-gc) = Normal ((KValue v, Value t), g') \ expr e2 e cd st g' = Normal ((v', t'), g'') \ g'' \ g'" unfolding gc_def by simp + ultimately show "g4 \ g" using lift_gas[OF lift] `\ g \ gc` by (simp split:result.split_asm Stackvalue.split_asm Type.split_asm prod.split_asm) + qed + qed +next + case (22 e1 e2 e cd st g) + define gc where "gc = costs\<^sub>e (EQUAL e1 e2) e cd st" + show ?case + proof (rule allI[THEN allI, OF impI]) + fix g4 xa assume e_def: "expr (EQUAL e1 e2) e cd st g = Normal (xa, g4)" + then show "g4 \ g" + proof (cases) + assume "g \ gc" + with 22(1) e_def show ?thesis using expr.psimps(14) gc_def by simp + next + assume "\ g \ gc" + then have *: "assert Gas ((<) (costs\<^sub>e (EQUAL e1 e2) e cd st)) g = Normal ((), g)" using gc_def by simp + with 22(1) e_def gc_def have lift:"lift expr equal e1 e2 e cd st (g - gc) = Normal (xa, g4)" using expr.psimps(14)[OF 22(1)] by simp + moreover have **: "modify (\g. g - costs\<^sub>e (EQUAL e1 e2) e cd st) g = Normal ((), g - costs\<^sub>e (EQUAL e1 e2) e cd st)" by simp + with 22(2)[OF * **] have "\g4' v4 t4. expr e1 e cd st (g-gc) = Normal ((v4, t4), g4') \ g4' \ g - gc" unfolding gc_def by simp + moreover obtain v g' where ***: "expr e1 e cd st (g - costs\<^sub>e (EQUAL e1 e2) e cd st) = Normal (v, g')" using expr.psimps(14)[OF 22(1)] e_def by (simp split:if_split_asm result.split_asm prod.split_asm) + with 22(3)[OF * ** ***] have "\v t g' v' t' g''. expr e1 e cd st (g-gc) = Normal ((KValue v, Value t), g') \ expr e2 e cd st g' = Normal ((v', t'), g'') \ g'' \ g'" unfolding gc_def by simp + ultimately show "g4 \ g" using lift_gas[OF lift] `\ g \ gc` by (simp split:result.split_asm Stackvalue.split_asm Type.split_asm prod.split_asm) + qed + qed +next + case (23 e1 e2 e cd st g) + define gc where "gc = costs\<^sub>e (AND e1 e2) e cd st" + show ?case + proof (rule allI[THEN allI, OF impI]) + fix g4 xa assume e_def: "expr (AND e1 e2) e cd st g = Normal (xa, g4)" + then show "g4 \ g" + proof (cases) + assume "g \ gc" + with 23(1) e_def show ?thesis using expr.psimps(15) gc_def by simp + next + assume "\ g \ gc" + then have *: "assert Gas ((<) (costs\<^sub>e (AND e1 e2) e cd st)) g = Normal ((), g)" using gc_def by simp + with 23(1) e_def gc_def have lift:"lift expr vtand e1 e2 e cd st (g - gc) = Normal (xa, g4)" using expr.psimps(15)[OF 23(1)] by simp + moreover have **: "modify (\g. g - costs\<^sub>e (AND e1 e2) e cd st) g = Normal ((), g - costs\<^sub>e (AND e1 e2) e cd st)" by simp + with 23(2)[OF * **] have "\g4' v4 t4. expr e1 e cd st (g-gc) = Normal ((v4, t4), g4') \ g4' \ g - gc" unfolding gc_def by simp + moreover obtain v g' where ***: "expr e1 e cd st (g - costs\<^sub>e (AND e1 e2) e cd st) = Normal (v, g')" using expr.psimps(15)[OF 23(1)] e_def by (simp split:if_split_asm result.split_asm prod.split_asm) + with 23(3)[OF * ** ***] have "\v t g' v' t' g''. expr e1 e cd st (g-gc) = Normal ((KValue v, Value t), g') \ expr e2 e cd st g' = Normal ((v', t'), g'') \ g'' \ g'" unfolding gc_def by simp + ultimately show "g4 \ g" using lift_gas[OF lift] `\ g \ gc` by (simp split:result.split_asm Stackvalue.split_asm Type.split_asm prod.split_asm) + qed + qed +next + case (24 e1 e2 e cd st g) + define gc where "gc = costs\<^sub>e (OR e1 e2) e cd st" + show ?case + proof (rule allI[THEN allI, OF impI]) + fix g4 xa assume e_def: "expr (OR e1 e2) e cd st g = Normal (xa, g4)" + then show "g4 \ g" + proof (cases) + assume "g \ gc" + with 24(1) e_def show ?thesis using expr.psimps(16) gc_def by simp + next + assume "\ g \ gc" + then have *: "assert Gas ((<) (costs\<^sub>e (OR e1 e2) e cd st)) g = Normal ((), g)" using gc_def by simp + with 24(1) e_def gc_def have lift:"lift expr vtor e1 e2 e cd st (g - gc) = Normal (xa, g4)" using expr.psimps(16)[OF 24(1)] by simp + moreover have **: "modify (\g. g - costs\<^sub>e (OR e1 e2) e cd st) g = Normal ((), g - costs\<^sub>e (OR e1 e2) e cd st)" by simp + with 24(2)[OF * **] have "\g4' v4 t4. expr e1 e cd st (g-gc) = Normal ((v4, t4), g4') \ g4' \ g - gc" unfolding gc_def by simp + moreover obtain v g' where ***: "expr e1 e cd st (g - costs\<^sub>e (OR e1 e2) e cd st) = Normal (v, g')" using expr.psimps(16)[OF 24(1)] e_def by (simp split:if_split_asm result.split_asm prod.split_asm) + with 24(3)[OF * ** ***] have "\v t g' v' t' g''. expr e1 e cd st (g-gc) = Normal ((KValue v, Value t), g') \ expr e2 e cd st g' = Normal ((v', t'), g'') \ g'' \ g'" unfolding gc_def by simp + ultimately show "g4 \ g" using lift_gas[OF lift] `\ g \ gc` by (simp split:result.split_asm Stackvalue.split_asm Type.split_asm prod.split_asm) + qed + qed +next + case (25 i e cd st g) + define gc where "gc = costs\<^sub>e (LVAL i) e cd st" + show ?case + proof (rule allI[THEN allI, OF impI]) + fix g4 xa xaa assume e_def: "expr (LVAL i) e cd st g = Normal (xa, g4)" + then show "g4 \ g" + proof (cases) + assume "g \ gc" + with 25(1) e_def show ?thesis using expr.psimps(17) gc_def by simp + next + assume "\ g \ gc" + then have *: "assert Gas ((<) (costs\<^sub>e (LVAL i) e cd st)) g = Normal ((), g)" using gc_def by simp + then have "expr (LVAL i) e cd st g = rexp i e cd st (g - gc)" using expr.psimps(17)[OF 25(1)] gc_def by simp + then have "rexp i e cd st (g - gc) = Normal (xa, g4)" using e_def by simp + moreover have **: "modify (\g. g - costs\<^sub>e (LVAL i) e cd st) g = Normal ((), g - costs\<^sub>e (LVAL i) e cd st)" by simp + ultimately show ?thesis using 25(2)[OF * **] unfolding gc_def by (simp split:result.split_asm Stackvalue.split_asm Type.split_asm prod.split_asm) + qed + qed +next + case (26 i xe e cd st g) + define gc where "gc = costs\<^sub>e (CALL i xe) e cd st" + show ?case + proof (rule allI[THEN allI, OF impI]) + fix g4' v4 assume a1: "expr (CALL i xe) e cd st g = Normal (v4, g4')" + then have *: "assert Gas ((<) (costs\<^sub>e (CALL i xe) e cd st)) g = Normal ((), g)" using gc_def using expr.psimps(18)[OF 26(1)] by (simp split:if_split_asm) + have **: "modify (\g. g - costs\<^sub>e (CALL i xe) e cd st) g = Normal ((), g - gc)" unfolding gc_def by simp + show "g4' \ g" + proof (cases) + assume "g \ gc" + with 26(1) gc_def a1 show ?thesis using expr.psimps(18) by simp + next + assume gcost: "\ g \ gc" + then show ?thesis + proof (cases v4) + case (Pair l4 t4) + then show ?thesis + proof (cases "ep $$ contract e") + case None + with 26(1) a1 gc_def gcost show ?thesis using expr.psimps(18) by simp + next + case (Some a) + then show ?thesis + proof (cases a) + case p2: (fields ct x0 x1) + then have 1: "option Err (\_. ep $$ contract e) (g - gc) = Normal ((ct, x0, x1), (g - gc))" using Some by simp + then show ?thesis + proof (cases "fmlookup ct i") + case None + with 26(1) a1 gc_def gcost Some p2 show ?thesis using expr.psimps(18) by simp + next + case s1: (Some a) + then show ?thesis + proof (cases a) + case (Function x1) + then show ?thesis + proof (cases x1) + case p1: (fields fp ext x) + then show ?thesis + proof (cases ext) + case True + with 26(1) a1 gc_def gcost Some p2 s1 Function p1 show ?thesis using expr.psimps(18)[of i xe e cd st] by (auto split:unit.split_asm) + next + case False + then have 2: "(case ct $$ i of None \ throw Err | Some (Function (fp, True, xa)) \ throw Err | Some (Function (fp, False, xa)) \ return (fp, xa) | Some _ \ throw Err) (g - gc) = Normal ((fp, x), (g - gc))" using s1 Function p1 by simp + define e' where "e' = ffold (init ct) (emptyEnv (address e) (contract e) (sender e) (svalue e)) (fmdom ct)" + then show ?thesis + proof (cases "load False fp xe e' emptyStore emptyStore (memory st) e cd st (g - gc)") + case s4: (n a g') + then show ?thesis + proof (cases a) + case f2: (fields e\<^sub>l cd\<^sub>l k\<^sub>l m\<^sub>l) + then show ?thesis + proof (cases "expr x e\<^sub>l cd\<^sub>l (st\stack:=k\<^sub>l, memory:=m\<^sub>l\) g'") + case n2: (n a g'') + then show ?thesis + proof (cases a) + case p3: (Pair sv tp) + with 26(1) a1 gc_def gcost Some p2 s1 Function p1 e'_def s4 f2 n2 False + have "expr (CALL i xe) e cd st g = Normal ((sv, tp), g'')" + using expr.psimps(18)[OF 26(1)] by simp + with a1 have "g4' \ g''" by simp + also from 26(3)[OF * ** 1 _ 2 _ _ s4] e'_def f2 n2 p3 gcost gc_def + have "\ \ g'" by auto + also from 26(2)[OF * ** 1 _ 2 _] False e'_def f2 s4 + have "\ \ g - gc" unfolding ffold_init_def gc_def by blast + finally show ?thesis by simp + qed + next + case (e _) + with 26(1) a1 gc_def gcost Some p2 s1 Function p1 e'_def s4 f2 False show ?thesis using expr.psimps(18)[of i xe e cd st] by (auto simp add:Let_def split:unit.split_asm) + qed + qed + next + case (e _) + with 26(1) a1 gc_def gcost Some p2 s1 Function p1 e'_def False show ?thesis using expr.psimps(18)[of i xe e cd st] by (auto split:unit.split_asm) + qed + qed + qed + next + case (Method _) + with 26(1) a1 gc_def gcost Some p2 s1 show ?thesis using expr.psimps(18) by simp + next + case (Var _) + with 26(1) a1 gc_def gcost Some p2 s1 show ?thesis using expr.psimps(18) by simp + qed + qed + qed + qed + qed + qed + qed +next + case (27 ad i xe e cd st g) + define gc where "gc = costs\<^sub>e (ECALL ad i xe) e cd st" + show ?case + proof (rule allI[THEN allI, OF impI]) + fix g4' v4 assume a1: "expr (ECALL ad i xe) e cd st g = Normal (v4, g4')" + show "g4' \ g" + proof (cases v4) + case (Pair l4 t4) + then show ?thesis + proof (cases) + assume "g \ gc" + with gc_def a1 show ?thesis using expr.psimps(19)[OF 27(1)] by simp + next + assume gcost: "\ g \ gc" + then have *: "assert Gas ((<) (costs\<^sub>e (ECALL ad i xe) e cd st)) g = Normal ((), g)" using gc_def using expr.psimps(19)[OF 27(1)] by (simp split:if_split_asm) + have **: "modify (\g. g - costs\<^sub>e (ECALL ad i xe) e cd st) g = Normal ((), g - gc)" unfolding gc_def by simp + then show ?thesis + proof (cases "expr ad e cd st (g-gc)") + case (n a0 g') + then show ?thesis + proof (cases a0) + case p0: (Pair a b) + then show ?thesis + proof (cases a) + case (KValue adv) + then show ?thesis + proof (cases b) + case (Value x1) + then show ?thesis + proof (cases x1) + case (TSInt x1) + with a1 gc_def gcost n p0 KValue Value show ?thesis using expr.psimps(19)[OF 27(1)] by simp + next + case (TUInt x2) + with a1 gc_def gcost n p0 KValue Value show ?thesis using expr.psimps(19)[OF 27(1)] by simp + next + case TBool + with a1 gc_def gcost n p0 KValue Value show ?thesis using expr.psimps(19)[OF 27(1)] by simp + next + case TAddr + then have 1: "(case a0 of (KValue adv, Value TAddr) \ return adv | (KValue adv, Value _) \ throw Err | (KValue adv, _) \ throw Err | (_, b) \ throw Err) g' = Normal (adv, g')" using p0 KValue Value by simp + then show ?thesis + proof (cases "adv = address e") + case True + with a1 gc_def gcost n p0 KValue Value TAddr show ?thesis using expr.psimps(19)[OF 27(1)] by simp + next + case False + then have 2: "assert Err (\_. adv \ address e) g' = Normal ((), g')" by simp + then show ?thesis + proof (cases "type (accounts st adv)") + case None + with a1 gc_def gcost n p0 KValue Value TAddr False show ?thesis using expr.psimps(19)[OF 27(1)] by simp + next + case (Some x2) + then show ?thesis + proof (cases x2) + case EOA + with a1 gc_def gcost n p0 KValue Value TAddr False Some show ?thesis using expr.psimps(19)[OF 27(1)] by simp + next + case c: (Contract c) + then have 3: "(case type (accounts st adv) of Some (Contract c) \ return c | _ \ throw Err) g' = Normal (c, g')" using Some by simp + then show ?thesis + proof (cases "ep $$ c") + case None + with a1 gc_def gcost n p0 KValue Value TAddr False Some c show ?thesis using expr.psimps(19)[OF 27(1)] by simp + next + case s0: (Some a) + then show ?thesis + proof (cases a) + case p1: (fields ct xa xb) + then have 4: "option Err (\_. ep $$ c) g' = Normal ((ct, xa, xb), g')" using Some s0 by simp + then show ?thesis + proof (cases "ct $$ i") + case None + with a1 gc_def gcost n p0 KValue Value TAddr Some p1 False c s0 show ?thesis using expr.psimps(19)[OF 27(1)] by simp + next + case s1: (Some a) + then show ?thesis + proof (cases a) + case (Function x1) + then show ?thesis + proof (cases x1) + case p2: (fields fp ext x) + then show ?thesis + proof (cases ext) + case True + then have 5: "(case ct $$ i of None \ throw Err | Some (Function (fp, True, xa)) \ return (fp, xa) | Some (Function (fp, False, xa)) \ throw Err | Some _ \ throw Err) g' = Normal ((fp, x), g')" using s1 Function p2 by simp + define e' where "e' = ffold (init ct) (emptyEnv adv c (address e) (ShowL\<^sub>n\<^sub>a\<^sub>t 0)) (fmdom ct)" + then show ?thesis + proof (cases "load True fp xe e' emptyStore emptyStore emptyStore e cd st g'") + case n1: (n a g'') + then show ?thesis + proof (cases a) + case f2: (fields e\<^sub>l cd\<^sub>l k\<^sub>l m\<^sub>l) + show ?thesis + proof (cases "expr x e\<^sub>l cd\<^sub>l (st\stack:=k\<^sub>l, memory:=m\<^sub>l\) g''") + case n2: (n a g''') + then show ?thesis + proof (cases a) + case p3: (Pair sv tp) + with a1 gc_def gcost n p2 KValue Value TAddr Some p1 s1 Function p0 e'_def n1 f2 n2 True False s0 c + have "expr (ECALL ad i xe) e cd st g = Normal ((sv, tp), g''')" + using expr.psimps(19)[OF 27(1)] by auto + with a1 have "g4' \ g'''" by auto + also from 27(4)[OF * ** n 1 2 3 4 _ 5 _ _ n1] p3 f2 e'_def n2 + have "\ \ g''" by simp + also from 27(3)[OF * ** n 1 2 3 4 _ 5, of "(xa, xb)" fp x e'] e'_def n1 f2 + have "\ \ g'" by auto + also from 27(2)[OF * **] n + have "\ \ g" by simp + finally show ?thesis by simp + qed + next + case (e _) + with a1 gc_def gcost n p0 KValue Value TAddr False Some p1 s1 Function p2 e'_def n1 f2 True s0 c show ?thesis using expr.psimps(19)[OF 27(1)] by simp + qed + qed + next + case (e _) + with a1 gc_def gcost n p0 KValue Value TAddr False Some p1 s1 Function p2 e'_def True s0 c show ?thesis using expr.psimps(19)[OF 27(1)] by simp + qed + next + case f: False + with a1 gc_def gcost n p0 KValue Value TAddr Some p1 s1 Function p2 False s0 c show ?thesis using expr.psimps(19)[OF 27(1)] by simp + qed + qed + next + case (Method _) + with a1 gc_def gcost n p0 KValue Value TAddr Some p1 s1 False s0 c show ?thesis using expr.psimps(19)[OF 27(1)] by simp + next + case (Var _) + with a1 gc_def gcost n p0 KValue Value TAddr Some p1 s1 False s0 c show ?thesis using expr.psimps(19)[OF 27(1)] by simp + qed + qed + qed + qed + qed + qed + qed + qed + next + case (Calldata _) + with a1 gc_def gcost n p0 KValue show ?thesis using expr.psimps(19)[OF 27(1)] by simp + next + case (Memory _) + with a1 gc_def gcost n p0 KValue show ?thesis using expr.psimps(19)[OF 27(1)] by simp + next + case (Storage _) + with a1 gc_def gcost n p0 KValue show ?thesis using expr.psimps(19)[OF 27(1)] by simp + qed + next + case (KCDptr _) + with a1 gc_def gcost n p0 show ?thesis using expr.psimps(19)[OF 27(1)] by simp + next + case (KMemptr _) + with a1 gc_def gcost n p0 show ?thesis using expr.psimps(19)[OF 27(1)] by simp + next + case (KStoptr _) + with a1 gc_def gcost n p0 show ?thesis using expr.psimps(19)[OF 27(1)] by simp + qed + qed + next + case (e _) + with a1 gc_def gcost show ?thesis using expr.psimps(19)[OF 27(1)] by simp + qed + qed + qed + qed +next + case (28 cp i\<^sub>p t\<^sub>p pl e el e\<^sub>v' cd' sck' mem' e\<^sub>v cd st g) + then show ?case + proof (cases "expr e e\<^sub>v cd st g") + case (n a g'') + then show ?thesis + proof (cases a) + case (Pair v t) + then show ?thesis + proof (cases "decl i\<^sub>p t\<^sub>p (Some (v,t)) cp cd (memory st) (storage st) (cd', mem', sck', e\<^sub>v')") + case None + with 28(1) n Pair show ?thesis using load.psimps(1) by simp + next + case (Some a) + show ?thesis + proof (cases a) + case (fields cd'' m'' k'' e\<^sub>v'') + then have 1: "(case decl i\<^sub>p t\<^sub>p (Some (v, t)) cp cd (memory st) (storage st) (cd', mem', sck', e\<^sub>v') of None \ throw Err | Some (c, m, k, e) \ return (c, m, k, e)) g'' = Normal ((cd'',m'',k'',e\<^sub>v''), g'')" using Some by simp + show ?thesis + proof ((rule allI)+,(rule impI)) + fix ev cda k m g' assume load_def: "load cp ((i\<^sub>p, t\<^sub>p) # pl) (e # el) e\<^sub>v' cd' sck' mem' e\<^sub>v cd st g = Normal ((ev, cda, k, m), g')" + with n Pair Some fields have "load cp ((i\<^sub>p, t\<^sub>p) # pl) (e # el) e\<^sub>v' cd' sck' mem' e\<^sub>v cd st g = load cp pl el e\<^sub>v'' cd'' k'' m'' e\<^sub>v cd st g''" using load.psimps(1)[OF 28(1)] by simp + with load_def have "load cp pl el e\<^sub>v'' cd'' k'' m'' e\<^sub>v cd st g'' = Normal ((ev, cda, k, m), g')" by simp + with Pair fields have "g' \ g'' \ address ev = address e\<^sub>v'' \ sender ev = sender e\<^sub>v'' \ svalue ev = svalue e\<^sub>v''" using 28(3)[OF n Pair 1, of cd'' _ m''] by simp + moreover from n have "g'' \ g" using 28(2) by simp + moreover from Some fields have "address e\<^sub>v'' = address e\<^sub>v' \ sender e\<^sub>v'' = sender e\<^sub>v' \ svalue e\<^sub>v'' = svalue e\<^sub>v'" using decl_env by auto + ultimately show "g' \ g \ address ev = address e\<^sub>v' \ sender ev = sender e\<^sub>v' \ svalue ev = svalue e\<^sub>v'" by auto + qed + qed + qed + qed + next + case (e _) + with 28(1) show ?thesis using load.psimps(1) by simp + qed +next + case 29 + then show ?case using load.psimps(2) by auto +next + case 30 + then show ?case using load.psimps(3) by auto +next + case 31 + then show ?case using load.psimps(4) by auto +next + case (32 i e cd st g) + show ?case + proof (rule allI[THEN allI, OF impI]) + fix xa xaa assume "rexp (L.Id i) e cd st g = Normal (xa, xaa)" + then show "xaa \ g" using 32(1) rexp.psimps(1) by (simp split: option.split_asm Denvalue.split_asm Stackvalue.split_asm prod.split_asm Type.split_asm STypes.split_asm) + qed +next + case (33 i r e cd st g) + show ?case + proof (rule allI[THEN allI,OF impI]) + fix xa xaa assume rexp_def: "rexp (Ref i r) e cd st g = Normal (xa, xaa)" + show "xaa \ g" + proof (cases "fmlookup (denvalue e) i") + case None + with 33(1) show ?thesis using rexp.psimps rexp_def by simp + next + case (Some a) + then show ?thesis + proof (cases a) + case (Pair tp b) + then show ?thesis + proof (cases b) + case (Stackloc l) + then show ?thesis + proof (cases "accessStore l (stack st)") + case None + with 33(1) Some Pair Stackloc show ?thesis using rexp.psimps(2) rexp_def by simp + next + case s1: (Some a) + then show ?thesis + proof (cases a) + case (KValue x1) + with 33(1) Some Pair Stackloc s1 show ?thesis using rexp.psimps(2) rexp_def by simp + next + case (KCDptr l') + with 33 Some Pair Stackloc s1 show ?thesis using rexp.psimps(2)[of i r e cd st] rexp_def by (simp split: option.split_asm Memoryvalue.split_asm MTypes.split_asm prod.split_asm Type.split_asm StateMonad.result.split_asm) + next + case (KMemptr x3) + with 33 Some Pair Stackloc s1 show ?thesis using rexp.psimps(2)[of i r e cd st] rexp_def by (simp split: option.split_asm Memoryvalue.split_asm MTypes.split_asm prod.split_asm Type.split_asm StateMonad.result.split_asm) + next + case (KStoptr x4) + with 33 Some Pair Stackloc s1 show ?thesis using rexp.psimps(2)[of i r e cd st] rexp_def by (simp split: option.split_asm STypes.split_asm prod.split_asm Type.split_asm StateMonad.result.split_asm) + qed + qed + next + case (Storeloc x2) + with 33 Some Pair show ?thesis using rexp.psimps rexp_def by (simp split: option.split_asm STypes.split_asm prod.split_asm Type.split_asm StateMonad.result.split_asm) + qed + qed + qed + qed +next + case (34 e cd st g) + then show ?case using expr.psimps(20) by (simp split:nat.split) +qed + +text \Now we can define the termination function\ + +fun mgas + where "mgas (Inr (Inr (Inr l))) = snd (snd (snd (snd l)))" (*rexp*) + | "mgas (Inr (Inr (Inl l))) = snd (snd (snd (snd (snd (snd (snd (snd (snd (snd l)))))))))" (*load*) + | "mgas (Inr (Inl l)) = snd (snd (snd (snd l)))" (*expr*) + | "mgas (Inl (Inr l)) = snd (snd (snd (snd (snd (snd l)))))" (*ssel*) + | "mgas (Inl (Inl l)) = snd (snd (snd (snd (snd (snd (snd l))))))" (*msel*) + +fun msize + where "msize (Inr (Inr (Inr l))) = size (fst l)" + | "msize (Inr (Inr (Inl l))) = size_list size (fst (snd (snd l)))" + | "msize (Inr (Inl l)) = size (fst l)" + | "msize (Inl (Inr l)) = size_list size (fst (snd (snd l)))" + | "msize (Inl (Inl l)) = size_list size (fst (snd (snd (snd l))))" + +method msel_ssel_expr_load_rexp_dom = + match premises in e: "expr _ _ _ _ _ = Normal (_,_)" and d[thin]: "msel_ssel_expr_load_rexp_dom (Inr (Inl _))" \ \insert msel_ssel_expr_load_rexp_dom_gas(3)[OF d e]\ | + match premises in l: "load _ _ _ _ _ _ _ _ _ _ _ = Normal (_,_)" and d[thin]: "msel_ssel_expr_load_rexp_dom (Inr (Inr (Inl _)))" \ \insert msel_ssel_expr_load_rexp_dom_gas(4)[OF d l, THEN conjunct1]\ + +method costs = + match premises in "costs\<^sub>e (CALL i xe) e cd st < _" for i xe and e::Environment and cd::CalldataT and st::State \ \insert call_not_zero[of (unchecked) i xe e cd st]\ | + match premises in "costs\<^sub>e (ECALL ad i xe) e cd st < _" for ad i xe and e::Environment and cd::CalldataT and st::State \ \insert ecall_not_zero[of (unchecked) ad i xe e cd st]\ + +termination msel + apply (relation "measures [mgas, msize]") + apply (auto split:Member.split_asm prod.split_asm bool.split_asm if_split_asm Stackvalue.split_asm option.split_asm Type.split_asm Types.split_asm Memoryvalue.split_asm atype.split_asm) + apply (msel_ssel_expr_load_rexp_dom+, costs?, arith)+ + done + +subsection \Gas Reduction\ + +text \ + The following corollary is a generalization of @{thm [source] msel_ssel_expr_load_rexp_dom_gas}. + We first prove that the function is defined for all input values and then obtain the final result as a corollary. +\ + +lemma msel_ssel_expr_load_rexp_dom: + "msel_ssel_expr_load_rexp_dom (Inl (Inl (c1, t1, l1, xe1, ev1, cd1, st1, g1)))" + "msel_ssel_expr_load_rexp_dom (Inl (Inr (t2, l2, xe2, ev2, cd2, st2, g2)))" + "msel_ssel_expr_load_rexp_dom (Inr (Inl (e4, ev4, cd4, st4, g4)))" + "msel_ssel_expr_load_rexp_dom (Inr (Inr (Inl (lcp, lis, lxs, lev0, lcd0, lk, lm, lev, lcd, lst, lg))))" + "msel_ssel_expr_load_rexp_dom (Inr (Inr (Inr (l3, ev3, cd3, st3, g3))))" +by (induct rule: msel_ssel_expr_load_rexp.induct +[where ?P1.0="\c1 t1 l1 xe1 ev1 cd1 st1 g1. msel_ssel_expr_load_rexp_dom (Inl (Inl (c1, t1, l1, xe1, ev1, cd1, st1, g1)))" + and ?P2.0="\t2 l2 xe2 ev2 cd2 st2 g2. msel_ssel_expr_load_rexp_dom (Inl (Inr (t2, l2, xe2, ev2, cd2, st2, g2)))" + and ?P3.0="\e4 ev4 cd4 st4 g4. msel_ssel_expr_load_rexp_dom (Inr (Inl (e4, ev4, cd4, st4, g4)))" + and ?P4.0="\lcp lis lxs lev0 lcd0 lk lm lev lcd lst lg. msel_ssel_expr_load_rexp_dom (Inr (Inr (Inl (lcp, lis, lxs, lev0, lcd0, lk, lm, lev, lcd, lst, lg))))" + and ?P5.0="\l3 ev3 cd3 st3 g3. msel_ssel_expr_load_rexp_dom (Inr (Inr (Inr (l3, ev3, cd3, st3, g3))))" +], simp_all add: msel_ssel_expr_load_rexp.domintros) + +lemmas msel_ssel_expr_load_rexp_gas = + msel_ssel_expr_load_rexp_dom_gas(1)[OF msel_ssel_expr_load_rexp_dom(1)] + msel_ssel_expr_load_rexp_dom_gas(2)[OF msel_ssel_expr_load_rexp_dom(2)] + msel_ssel_expr_load_rexp_dom_gas(3)[OF msel_ssel_expr_load_rexp_dom(3)] + msel_ssel_expr_load_rexp_dom_gas(4)[OF msel_ssel_expr_load_rexp_dom(4)] + msel_ssel_expr_load_rexp_dom_gas(5)[OF msel_ssel_expr_load_rexp_dom(5)] + +lemma expr_sender: + assumes "expr SENDER e cd st g = Normal ((KValue adv, Value TAddr), g')" + shows "adv = sender e" using assms by (simp split:if_split_asm) + +declare expr.simps[simp del] +declare load.simps[simp del] +declare ssel.simps[simp del] +declare msel.simps[simp del] +declare rexp.simps[simp del] + +end + +end diff --git a/thys/Solidity/ROOT b/thys/Solidity/ROOT --- a/thys/Solidity/ROOT +++ b/thys/Solidity/ROOT @@ -1,27 +1,32 @@ -chapter AFP - -session Solidity = "HOL-Library" + - options [timeout = 7200] - sessions - "HOL-Eisbach" - theories - ReadShow - StateMonad - Valuetypes - Storage - Accounts - Environment - Statements - Solidity_Main - Solidity_Symbex - Solidity_Evaluator - Constant_Folding - Reentrancy - theories [condition = ISABELLE_GHC] - Compile_Evaluator - document_files - "root.tex" - "root.bib" - "orcidlink.sty" - export_files (in ".") [2] "*:**.hs" "*:**.ML" - export_files (in "solidity-evaluator/bin") [1] "*:solidity-evaluator" +chapter AFP + +session "Solidity" = "HOL-Library" + + options [timeout = 7200] + sessions + "HOL-Eisbach" + "Finite-Map-Extras" + theories + Utils + ReadShow + StateMonad + Valuetypes + Storage + Accounts + Environment + Contracts + Expressions + Statements + Solidity_Main + Solidity_Symbex + Solidity_Evaluator + Weakest_Precondition + Reentrancy + Constant_Folding + theories [condition = ISABELLE_GHC] + Compile_Evaluator + document_files + "root.tex" + "root.bib" + "orcidlink.sty" + export_files (in ".") [2] "*:**.hs" "*:**.ML" + export_files (in "solidity-evaluator/bin") [1] "*:solidity-evaluator" diff --git a/thys/Solidity/ReadShow.thy b/thys/Solidity/ReadShow.thy --- a/thys/Solidity/ReadShow.thy +++ b/thys/Solidity/ReadShow.thy @@ -1,438 +1,457 @@ -section\Converting Types to Strings and Back Again\ -theory ReadShow - imports - Solidity_Symbex -begin - -text\ - In the following, we formalize a family of projection (and injection) functions for injecting - (projecting) basic types (i.e., @{type \nat\}, @{type \int\}, and @{type \bool\} in (out) of the - domains of strings. We provide variants for the two string representations of Isabelle/HOL, namely - @{type \string\} and @{type \String.literal\}. -\ - - -subsubsection\Bool\ -definition - \Read\<^sub>b\<^sub>o\<^sub>o\<^sub>l s = (if s = ''True'' then True else False)\ -definition - \Show\<^sub>b\<^sub>o\<^sub>o\<^sub>l b = (if b then ''True'' else ''False'')\ -definition - \STR_is_bool s = (Show\<^sub>b\<^sub>o\<^sub>o\<^sub>l (Read\<^sub>b\<^sub>o\<^sub>o\<^sub>l s) = s)\ - -declare Read\<^sub>b\<^sub>o\<^sub>o\<^sub>l_def [solidity_symbex] - Show\<^sub>b\<^sub>o\<^sub>o\<^sub>l_def [solidity_symbex] - -lemma Show_Read_bool_id: \STR_is_bool s \ (Show\<^sub>b\<^sub>o\<^sub>o\<^sub>l (Read\<^sub>b\<^sub>o\<^sub>o\<^sub>l s) = s)\ - using STR_is_bool_def by fastforce - -lemma STR_is_bool_split: \STR_is_bool s \ s = ''False'' \ s = ''True''\ - by(auto simp: STR_is_bool_def Read\<^sub>b\<^sub>o\<^sub>o\<^sub>l_def Show\<^sub>b\<^sub>o\<^sub>o\<^sub>l_def) - -lemma Read_Show_bool_id: \Read\<^sub>b\<^sub>o\<^sub>o\<^sub>l (Show\<^sub>b\<^sub>o\<^sub>o\<^sub>l b) = b\ - by(auto simp: Read\<^sub>b\<^sub>o\<^sub>o\<^sub>l_def Show\<^sub>b\<^sub>o\<^sub>o\<^sub>l_def) - -definition ReadL\<^sub>b\<^sub>o\<^sub>o\<^sub>l::\String.literal \ bool\ (\\_\\) where - \ReadL\<^sub>b\<^sub>o\<^sub>o\<^sub>l s = (if s = STR ''True'' then True else False)\ -definition ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l:: \bool \ String.literal\ (\\_\\) where - \ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l b = (if b then STR ''True'' else STR ''False'')\ -definition - \strL_is_bool' s = (ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l (ReadL\<^sub>b\<^sub>o\<^sub>o\<^sub>l s) = s)\ - -declare ReadL\<^sub>b\<^sub>o\<^sub>o\<^sub>l_def [solidity_symbex] - ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l_def [solidity_symbex] - - -lemma Show_Read_bool'_id: \strL_is_bool' s \ (ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l (ReadL\<^sub>b\<^sub>o\<^sub>o\<^sub>l s) = s)\ - using strL_is_bool'_def by fastforce - -lemma strL_is_bool'_split: \strL_is_bool' s \ s = STR ''False'' \ s = STR ''True''\ - by(auto simp: strL_is_bool'_def ReadL\<^sub>b\<^sub>o\<^sub>o\<^sub>l_def ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l_def) - -lemma Read_Show_bool'_id: \ReadL\<^sub>b\<^sub>o\<^sub>o\<^sub>l (ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l b) = b\ - by(auto simp: ReadL\<^sub>b\<^sub>o\<^sub>o\<^sub>l_def ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l_def) - - -subsubsection\Natural Numbers\ - -definition nat_of_digit :: \char \ nat\ where - \nat_of_digit c = - (if c = CHR ''0'' then 0 - else if c = CHR ''1'' then 1 - else if c = CHR ''2'' then 2 - else if c = CHR ''3'' then 3 - else if c = CHR ''4'' then 4 - else if c = CHR ''5'' then 5 - else if c = CHR ''6'' then 6 - else if c = CHR ''7'' then 7 - else if c = CHR ''8'' then 8 - else if c = CHR ''9'' then 9 - else undefined)\ - -declare nat_of_digit_def [solidity_symbex] - -definition digit_of_nat :: \nat \ char\ where - \digit_of_nat x = - (if x = 0 then CHR ''0'' - else if x = 1 then CHR ''1'' - else if x = 2 then CHR ''2'' - else if x = 3 then CHR ''3'' - else if x = 4 then CHR ''4'' - else if x = 5 then CHR ''5'' - else if x = 6 then CHR ''6'' - else if x = 7 then CHR ''7'' - else if x = 8 then CHR ''8'' - else if x = 9 then CHR ''9'' - else undefined)\ - -declare digit_of_nat_def [solidity_symbex] - -lemma nat_of_digit_digit_of_nat_id: - \x < 10 \ nat_of_digit (digit_of_nat x) = x\ - by(simp add:nat_of_digit_def digit_of_nat_def) - -lemma img_digit_of_nat: -\n < 10 \ digit_of_nat n \ {CHR ''0'', CHR ''1'', CHR ''2'', CHR ''3'', CHR ''4'', - CHR ''5'', CHR ''6'', CHR ''7'', CHR ''8'', CHR ''9''}\ - by(simp add:nat_of_digit_def digit_of_nat_def) - -lemma digit_of_nat_nat_of_digit_id: - \c \ {CHR ''0'', CHR ''1'', CHR ''2'', CHR ''3'', CHR ''4'', - CHR ''5'', CHR ''6'', CHR ''7'', CHR ''8'', CHR ''9''} - \ digit_of_nat (nat_of_digit c) = c\ - by(code_simp, auto) - -definition - nat_implode :: \'a::{numeral,power,zero} list \ 'a\ where - \nat_implode n = foldr (+) (map (\ (p,d) \ 10 ^ p * d) (enumerate 0 (rev n))) 0\ - -declare nat_implode_def [solidity_symbex] - -fun nat_explode' :: \nat \ nat list\ where - \nat_explode' x = (case x < 10 of True \ [x mod 10] - | _ \ (x mod 10 )#(nat_explode' (x div 10)))\ - -definition - nat_explode :: \nat \ nat list\ where - \nat_explode x = (rev (nat_explode' x))\ - -declare nat_explode_def [solidity_symbex] - -lemma nat_explode'_not_empty: \nat_explode' n \ []\ - by (smt (z3) list.simps(3) nat_explode'.simps) - -lemma nat_explode_not_empty: \nat_explode n \ []\ - using nat_explode'_not_empty nat_explode_def by auto - -lemma nat_explode'_ne_suc: \\ n. nat_explode' (Suc n) \ nat_explode' n\ - by(rule exI [of _ \0::nat\], simp) - -lemma nat_explode'_digit: \hd (nat_explode' n ) < 10\ -proof(induct \n\) - case 0 - then show ?case by simp -next - case (Suc n) - then show ?case proof (cases \n < 9\) - case True - then show ?thesis by simp - next - case False - then show ?thesis - by simp - qed -qed - -lemma div_ten_less: \n \ 0 \ ((n::nat) div 10) < n\ - by simp - -lemma unroll_nat_explode': - \\ n < 10 \ (case n < 10 of True \ [n mod 10] | False \ n mod 10 # nat_explode' (n div 10)) = - (n mod 10 # nat_explode' (n div 10))\ - by simp - -lemma nat_explode_mod_10_ident: \map (\ x. x mod 10) (nat_explode' n) = nat_explode' n\ -proof (cases \n < 10\) - case True - then show ?thesis by simp -next - case False - then show ?thesis - proof (induct \n\ rule: nat_less_induct) - case (1 n) note * = this - then show ?case - using div_ten_less apply(simp (no_asm)) - using unroll_nat_explode'[of n] * - by (smt (z3) list.simps(8) list.simps(9) mod_div_trivial mod_eq_self_iff_div_eq_0 - nat_explode'.simps zero_less_numeral) - qed -qed - -lemma nat_explode'_digits: - \\d \ set (nat_explode' n). d < 10\ -proof - fix d - assume \d \ set (nat_explode' n)\ - then have \d \ set (map (\m. m mod 10) (nat_explode' n))\ - by (simp only: nat_explode_mod_10_ident) - then show \d < 10\ - by auto -qed - -lemma nat_explode_digits: - \\d \ set (nat_explode n). d < 10\ - using nat_explode'_digits [of n] by (simp only: nat_explode_def set_rev) - -value \nat_implode(nat_explode 42) = 42\ -value \nat_explode (Suc 21)\ - - -lemma nat_implode_append: - \nat_implode (a@[b]) = (1*b + foldr (+) (map (\(p, y). 10 ^ p * y) (enumerate (Suc 0) (rev a))) 0 )\ - by(simp add:nat_implode_def) - -lemma enumerate_suc: \enumerate (Suc n) l = map (\ (a,b). (a+1::nat,b)) (enumerate n l)\ -proof (induction \l\) - case Nil - then show ?case by simp -next - case (Cons a x) note * = this - then show ?case apply(simp only:enumerate_simps) - - apply(simp only:\enumerate (Suc n) x = map (\a. case a of (a, b) \ (a + 1, b)) (enumerate n x)\[symmetric]) - apply(simp) - using * - by (metis apfst_conv cond_case_prod_eta enumerate_Suc_eq) -qed - -lemma mult_assoc_aux1: - \(\(p, y). 10 ^ p * y) \ (\(a, y). (Suc a, y)) = (\(p, y). (10::nat) * (10 ^ p) * y)\ - by(auto simp:o_def) - -lemma fold_map_transfer: - \(foldr (+) (map (\(x,y). 10 * (f (x,y))) l) (0::nat)) = 10 * (foldr (+) (map (\x. (f x)) l) (0::nat))\ -proof(induct \l\) - case Nil - then show ?case by(simp) -next - case (Cons a l) - then show ?case by(simp) -qed - -lemma mult_assoc_aux2: \(\(p, y). 10 * 10 ^ p * (y::nat)) = (\(p, y). 10 * (10 ^ p * y))\ - by(auto) - -lemma nat_implode_explode_id: \nat_implode (nat_explode n) = n\ -proof (cases \n=0\) - case True note * = this - then show ?thesis - by (simp add: nat_explode_def nat_implode_def) -next - case False - then show ?thesis - proof (induct \n\ rule: nat_less_induct) - case (1 n) note * = this - then have - **: \nat_implode (nat_explode (n div 10)) = n div 10\ - proof (cases \n div 10 = 0\) - case True - then show ?thesis by(simp add: nat_explode_def nat_implode_def) - next - case False - then show ?thesis - using div_ten_less[of \n\] * - by(simp) - qed - then show ?case - proof (cases \n < 10\) - case True - then show ?thesis by(simp add: nat_explode_def nat_implode_def) - next - case False note *** = this - then show ?thesis - apply(simp (no_asm) add:nat_explode_def del:rev_rev_ident) - apply(simp only: bool.case(2)) - apply(simp del:nat_explode'.simps rev_rev_ident) - apply(fold nat_explode_def) - apply(simp only:nat_implode_append) - apply(simp add:enumerate_suc) - apply(simp only:mult_assoc_aux1) - using mult_assoc_aux2 apply(simp) - using fold_map_transfer[of \\(p, y). 10 ^ p * y\ - \(enumerate 0 (rev (nat_explode (n div 10))))\, simplified] - apply(simp) apply(fold nat_implode_def) using ** - by simp - qed - qed -qed - -definition - Read\<^sub>n\<^sub>a\<^sub>t :: \string \ nat\ where - \Read\<^sub>n\<^sub>a\<^sub>t s = nat_implode (map nat_of_digit s)\ - -definition - Show\<^sub>n\<^sub>a\<^sub>t::"nat \ string" where - \Show\<^sub>n\<^sub>a\<^sub>t n = map digit_of_nat (nat_explode n)\ - -declare Read\<^sub>n\<^sub>a\<^sub>t_def [solidity_symbex] - Show\<^sub>n\<^sub>a\<^sub>t_def [solidity_symbex] - -definition - \STR_is_nat s = (Show\<^sub>n\<^sub>a\<^sub>t (Read\<^sub>n\<^sub>a\<^sub>t s) = s)\ - -value \Read\<^sub>n\<^sub>a\<^sub>t ''10''\ -value \Show\<^sub>n\<^sub>a\<^sub>t 10\ -value \Read\<^sub>n\<^sub>a\<^sub>t (Show\<^sub>n\<^sub>a\<^sub>t (10)) = 10\ -value \Show\<^sub>n\<^sub>a\<^sub>t (Read\<^sub>n\<^sub>a\<^sub>t (''10'')) = ''10''\ - -lemma Show_nat_not_neg: - \set (Show\<^sub>n\<^sub>a\<^sub>t n) \{CHR ''0'', CHR ''1'', CHR ''2'', CHR ''3'', CHR ''4'', - CHR ''5'', CHR ''6'', CHR ''7'', CHR ''8'', CHR ''9''}\ - unfolding Show\<^sub>n\<^sub>a\<^sub>t_def - using nat_explode_digits[of n] img_digit_of_nat imageE image_set subsetI - by (smt (verit) imageE image_set subsetI) - -lemma Show_nat_not_empty: \(Show\<^sub>n\<^sub>a\<^sub>t n) \ []\ - by (simp add: Show\<^sub>n\<^sub>a\<^sub>t_def nat_explode_not_empty) - -lemma not_hd: \L \ [] \ e \ set(L) \ hd L \ e\ - by auto - -lemma Show_nat_not_neg'': \hd (Show\<^sub>n\<^sub>a\<^sub>t n) \ (CHR ''-'')\ - using Show_nat_not_neg[of \n\] - Show_nat_not_empty[of \n\] not_hd[of \Show\<^sub>n\<^sub>a\<^sub>t n\] - by auto - -lemma Show_Read_nat_id: \STR_is_nat s \ (Show\<^sub>n\<^sub>a\<^sub>t (Read\<^sub>n\<^sub>a\<^sub>t s) = s)\ - by(simp add:STR_is_nat_def) - -lemma bar': \\ d \ set l . d < 10 \ map nat_of_digit (map digit_of_nat l) = l\ - using nat_of_digit_digit_of_nat_id - by (simp add: map_idI) - -lemma Read_Show_nat_id: \Read\<^sub>n\<^sub>a\<^sub>t(Show\<^sub>n\<^sub>a\<^sub>t n) = n\ - apply(unfold Read\<^sub>n\<^sub>a\<^sub>t_def Show\<^sub>n\<^sub>a\<^sub>t_def) - using bar' nat_of_digit_digit_of_nat_id nat_explode_digits - using nat_implode_explode_id - by presburger - -definition - ReadL\<^sub>n\<^sub>a\<^sub>t :: \String.literal \ nat\ (\\_\\) where - \ReadL\<^sub>n\<^sub>a\<^sub>t = Read\<^sub>n\<^sub>a\<^sub>t \ String.explode\ - -definition - ShowL\<^sub>n\<^sub>a\<^sub>t::\nat \ String.literal\ (\\_\\)where - \ShowL\<^sub>n\<^sub>a\<^sub>t = String.implode \ Show\<^sub>n\<^sub>a\<^sub>t\ - -declare ReadL\<^sub>n\<^sub>a\<^sub>t_def [solidity_symbex] - ShowL\<^sub>n\<^sub>a\<^sub>t_def [solidity_symbex] - - -definition - \strL_is_nat' s = (ShowL\<^sub>n\<^sub>a\<^sub>t (ReadL\<^sub>n\<^sub>a\<^sub>t s) = s)\ - -value \\STR ''10''\::nat\ -value \ReadL\<^sub>n\<^sub>a\<^sub>t (STR ''10'')\ -value \\10::nat\\ -value \ShowL\<^sub>n\<^sub>a\<^sub>t 10\ -value \ReadL\<^sub>n\<^sub>a\<^sub>t (ShowL\<^sub>n\<^sub>a\<^sub>t (10)) = 10\ -value \ShowL\<^sub>n\<^sub>a\<^sub>t (ReadL\<^sub>n\<^sub>a\<^sub>t (STR ''10'')) = STR ''10''\ - -lemma Show_Read_nat'_id: \strL_is_nat' s \ (ShowL\<^sub>n\<^sub>a\<^sub>t (ReadL\<^sub>n\<^sub>a\<^sub>t s) = s)\ - by(simp add:strL_is_nat'_def) - -lemma digits_are_ascii: - \c \ {CHR ''0'', CHR ''1'', CHR ''2'', CHR ''3'', CHR ''4'', - CHR ''5'', CHR ''6'', CHR ''7'', CHR ''8'', CHR ''9''} - \ String.ascii_of c = c\ - by auto - -lemma Show\<^sub>n\<^sub>a\<^sub>t_ascii: \map String.ascii_of (Show\<^sub>n\<^sub>a\<^sub>t n) = Show\<^sub>n\<^sub>a\<^sub>t n\ - using Show_nat_not_neg digits_are_ascii - by (meson map_idI subsetD) - - -lemma Read_Show_nat'_id: \ReadL\<^sub>n\<^sub>a\<^sub>t(ShowL\<^sub>n\<^sub>a\<^sub>t n) = n\ - apply(unfold ReadL\<^sub>n\<^sub>a\<^sub>t_def ShowL\<^sub>n\<^sub>a\<^sub>t_def, simp) - by (simp add: Show\<^sub>n\<^sub>a\<^sub>t_ascii Read_Show_nat_id) - - -subsubsection\Integer\ -definition - Read\<^sub>i\<^sub>n\<^sub>t :: \string \ int\ where - \Read\<^sub>i\<^sub>n\<^sub>t x = (if hd x = (CHR ''-'') then -(int (Read\<^sub>n\<^sub>a\<^sub>t (tl x))) else int (Read\<^sub>n\<^sub>a\<^sub>t x))\ - -definition - Show\<^sub>i\<^sub>n\<^sub>t::\int \ string\ where - \Show\<^sub>i\<^sub>n\<^sub>t i = (if i < 0 then (CHR ''-'')#(Show\<^sub>n\<^sub>a\<^sub>t (nat (-i))) - else Show\<^sub>n\<^sub>a\<^sub>t (nat i))\ - -definition - \STR_is_int s = (Show\<^sub>i\<^sub>n\<^sub>t (Read\<^sub>i\<^sub>n\<^sub>t s) = s)\ - - -declare Read\<^sub>i\<^sub>n\<^sub>t_def [solidity_symbex] - Show\<^sub>i\<^sub>n\<^sub>t_def [solidity_symbex] - -value \Read\<^sub>i\<^sub>n\<^sub>t (Show\<^sub>i\<^sub>n\<^sub>t 10) = 10\ -value \Read\<^sub>i\<^sub>n\<^sub>t (Show\<^sub>i\<^sub>n\<^sub>t (-10)) = -10\ - -value \Show\<^sub>i\<^sub>n\<^sub>t (Read\<^sub>i\<^sub>n\<^sub>t (''10'')) = ''10''\ -value \Show\<^sub>i\<^sub>n\<^sub>t (Read\<^sub>i\<^sub>n\<^sub>t (''-10'')) = ''-10''\ - -lemma Show_Read_id: \STR_is_int s \ (Show\<^sub>i\<^sub>n\<^sub>t (Read\<^sub>i\<^sub>n\<^sub>t s) = s)\ - by(simp add:STR_is_int_def) - -lemma Read_Show_id: \Read\<^sub>i\<^sub>n\<^sub>t(Show\<^sub>i\<^sub>n\<^sub>t(x)) = x\ - apply(code_simp) - apply(auto simp:Show_nat_not_neg Read_Show_nat_id)[1] - apply(thin_tac \\ x < 0 \) - using Show_nat_not_neg'' - by blast - -lemma STR_is_int_Show: \STR_is_int (Show\<^sub>i\<^sub>n\<^sub>t n)\ - by(simp add:STR_is_int_def Read_Show_id) - -definition - ReadL\<^sub>i\<^sub>n\<^sub>t :: \String.literal \ int\ (\\_\\) where - \ReadL\<^sub>i\<^sub>n\<^sub>t = Read\<^sub>i\<^sub>n\<^sub>t \ String.explode\ - -definition - ShowL\<^sub>i\<^sub>n\<^sub>t::\int \ String.literal\ (\\_\\) where - \ShowL\<^sub>i\<^sub>n\<^sub>t =String.implode \ Show\<^sub>i\<^sub>n\<^sub>t\ - -definition - \strL_is_int' s = (ShowL\<^sub>i\<^sub>n\<^sub>t (ReadL\<^sub>i\<^sub>n\<^sub>t s) = s)\ - -declare ReadL\<^sub>i\<^sub>n\<^sub>t_def [solidity_symbex] - ShowL\<^sub>i\<^sub>n\<^sub>t_def [solidity_symbex] - -value \ReadL\<^sub>i\<^sub>n\<^sub>t (ShowL\<^sub>i\<^sub>n\<^sub>t 10) = 10\ -value \ReadL\<^sub>i\<^sub>n\<^sub>t (ShowL\<^sub>i\<^sub>n\<^sub>t (-10)) = -10\ - -value \ShowL\<^sub>i\<^sub>n\<^sub>t (ReadL\<^sub>i\<^sub>n\<^sub>t (STR ''10'')) = STR ''10''\ -value \ShowL\<^sub>i\<^sub>n\<^sub>t (ReadL\<^sub>i\<^sub>n\<^sub>t (STR ''-10'')) = STR ''-10''\ - -lemma Show_ReadL_id: \strL_is_int' s \ (ShowL\<^sub>i\<^sub>n\<^sub>t (ReadL\<^sub>i\<^sub>n\<^sub>t s) = s)\ - by(simp add:strL_is_int'_def) - -lemma Read_ShowL_id: \ReadL\<^sub>i\<^sub>n\<^sub>t (ShowL\<^sub>i\<^sub>n\<^sub>t x) = x\ -proof(cases \x < 0\) - case True - then show ?thesis using ShowL\<^sub>i\<^sub>n\<^sub>t_def ReadL\<^sub>i\<^sub>n\<^sub>t_def Show\<^sub>i\<^sub>n\<^sub>t_def Show\<^sub>n\<^sub>a\<^sub>t_ascii - by (metis (no_types, lifting) Read_Show_id String.ascii_of_Char comp_def implode.rep_eq list.simps(9)) -next - case False - then show ?thesis using ShowL\<^sub>i\<^sub>n\<^sub>t_def ReadL\<^sub>i\<^sub>n\<^sub>t_def Show\<^sub>i\<^sub>n\<^sub>t_def Show\<^sub>n\<^sub>a\<^sub>t_ascii - by (metis Read_Show_id String.explode_implode_eq comp_apply) -qed - -lemma STR_is_int_ShowL: \strL_is_int' (ShowL\<^sub>i\<^sub>n\<^sub>t n)\ - by(simp add:strL_is_int'_def Read_ShowL_id) - -lemma String_Cancel: "a + (c::String.literal) = b + c \ a = b" -using String.plus_literal.rep_eq -by (metis append_same_eq literal.explode_inject) - -end - +section\Converting Types to Strings and Back Again\ +theory ReadShow + imports + Solidity_Symbex +begin + +text\ + In the following, we formalize a family of projection (and injection) functions for injecting + (projecting) basic types (i.e., @{type \nat\}, @{type \int\}, and @{type \bool\} in (out) of the + domains of strings. We provide variants for the two string representations of Isabelle/HOL, namely + @{type \string\} and @{type \String.literal\}. +\ + + +subsubsection\Bool\ +definition + \Read\<^sub>b\<^sub>o\<^sub>o\<^sub>l s = (if s = ''True'' then True else False)\ +definition + \Show\<^sub>b\<^sub>o\<^sub>o\<^sub>l b = (if b then ''True'' else ''False'')\ +definition + \STR_is_bool s = (Show\<^sub>b\<^sub>o\<^sub>o\<^sub>l (Read\<^sub>b\<^sub>o\<^sub>o\<^sub>l s) = s)\ + +declare Read\<^sub>b\<^sub>o\<^sub>o\<^sub>l_def [solidity_symbex] + Show\<^sub>b\<^sub>o\<^sub>o\<^sub>l_def [solidity_symbex] + +lemma Show_Read_bool_id: \STR_is_bool s \ (Show\<^sub>b\<^sub>o\<^sub>o\<^sub>l (Read\<^sub>b\<^sub>o\<^sub>o\<^sub>l s) = s)\ + using STR_is_bool_def by fastforce + +lemma STR_is_bool_split: \STR_is_bool s \ s = ''False'' \ s = ''True''\ + by(auto simp: STR_is_bool_def Read\<^sub>b\<^sub>o\<^sub>o\<^sub>l_def Show\<^sub>b\<^sub>o\<^sub>o\<^sub>l_def) + +lemma Read_Show_bool_id: \Read\<^sub>b\<^sub>o\<^sub>o\<^sub>l (Show\<^sub>b\<^sub>o\<^sub>o\<^sub>l b) = b\ + by(auto simp: Read\<^sub>b\<^sub>o\<^sub>o\<^sub>l_def Show\<^sub>b\<^sub>o\<^sub>o\<^sub>l_def) + +definition ReadL\<^sub>b\<^sub>o\<^sub>o\<^sub>l::\String.literal \ bool\ (\\_\\) where + \ReadL\<^sub>b\<^sub>o\<^sub>o\<^sub>l s = (if s = STR ''True'' then True else False)\ +definition ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l:: \bool \ String.literal\ (\\_\\) where + \ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l b = (if b then STR ''True'' else STR ''False'')\ +definition + \strL_is_bool' s = (ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l (ReadL\<^sub>b\<^sub>o\<^sub>o\<^sub>l s) = s)\ + +declare ReadL\<^sub>b\<^sub>o\<^sub>o\<^sub>l_def [solidity_symbex] + ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l_def [solidity_symbex] + + +lemma Show_Read_bool'_id: \strL_is_bool' s \ (ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l (ReadL\<^sub>b\<^sub>o\<^sub>o\<^sub>l s) = s)\ + using strL_is_bool'_def by fastforce + +lemma strL_is_bool'_split: \strL_is_bool' s \ s = STR ''False'' \ s = STR ''True''\ + by(auto simp: strL_is_bool'_def ReadL\<^sub>b\<^sub>o\<^sub>o\<^sub>l_def ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l_def) + +lemma Read_Show_bool'_id[simp]: \ReadL\<^sub>b\<^sub>o\<^sub>o\<^sub>l (ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l b) = b\ + by(auto simp: ReadL\<^sub>b\<^sub>o\<^sub>o\<^sub>l_def ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l_def) + +lemma true_neq_false[simp]: "ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True \ ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l False" + by (metis Read_Show_bool'_id) + +subsubsection\Natural Numbers\ + +definition nat_of_digit :: \char \ nat\ where + \nat_of_digit c = + (if c = CHR ''0'' then 0 + else if c = CHR ''1'' then 1 + else if c = CHR ''2'' then 2 + else if c = CHR ''3'' then 3 + else if c = CHR ''4'' then 4 + else if c = CHR ''5'' then 5 + else if c = CHR ''6'' then 6 + else if c = CHR ''7'' then 7 + else if c = CHR ''8'' then 8 + else if c = CHR ''9'' then 9 + else undefined)\ + +declare nat_of_digit_def [solidity_symbex] + +definition is_digit :: \char \ bool\ where + \is_digit c = + (if c = CHR ''0'' then True + else if c = CHR ''1'' then True + else if c = CHR ''2'' then True + else if c = CHR ''3'' then True + else if c = CHR ''4'' then True + else if c = CHR ''5'' then True + else if c = CHR ''6'' then True + else if c = CHR ''7'' then True + else if c = CHR ''8'' then True + else if c = CHR ''9'' then True + else if c = CHR ''-'' then True + else False)\ + + + +definition digit_of_nat :: \nat \ char\ where + \digit_of_nat x = + (if x = 0 then CHR ''0'' + else if x = 1 then CHR ''1'' + else if x = 2 then CHR ''2'' + else if x = 3 then CHR ''3'' + else if x = 4 then CHR ''4'' + else if x = 5 then CHR ''5'' + else if x = 6 then CHR ''6'' + else if x = 7 then CHR ''7'' + else if x = 8 then CHR ''8'' + else if x = 9 then CHR ''9'' + else undefined)\ + +declare digit_of_nat_def [solidity_symbex] + +lemma nat_of_digit_digit_of_nat_id: + \x < 10 \ nat_of_digit (digit_of_nat x) = x\ + by(simp add:nat_of_digit_def digit_of_nat_def) + +lemma img_digit_of_nat: +\n < 10 \ digit_of_nat n \ {CHR ''0'', CHR ''1'', CHR ''2'', CHR ''3'', CHR ''4'', + CHR ''5'', CHR ''6'', CHR ''7'', CHR ''8'', CHR ''9''}\ + by(simp add:nat_of_digit_def digit_of_nat_def) + +lemma digit_of_nat_nat_of_digit_id: + \c \ {CHR ''0'', CHR ''1'', CHR ''2'', CHR ''3'', CHR ''4'', + CHR ''5'', CHR ''6'', CHR ''7'', CHR ''8'', CHR ''9''} + \ digit_of_nat (nat_of_digit c) = c\ + by(code_simp, auto) + +definition + nat_implode :: \'a::{numeral,power,zero} list \ 'a\ where + \nat_implode n = foldr (+) (map (\ (p,d) \ 10 ^ p * d) (enumerate 0 (rev n))) 0\ + +declare nat_implode_def [solidity_symbex] + +fun nat_explode' :: \nat \ nat list\ where + \nat_explode' x = (case x < 10 of True \ [x mod 10] + | _ \ (x mod 10 )#(nat_explode' (x div 10)))\ + +definition + nat_explode :: \nat \ nat list\ where + \nat_explode x = (rev (nat_explode' x))\ + +declare nat_explode_def [solidity_symbex] + +lemma nat_explode'_not_empty: \nat_explode' n \ []\ + by (smt (z3) list.simps(3) nat_explode'.simps) + +lemma nat_explode_not_empty: \nat_explode n \ []\ + using nat_explode'_not_empty nat_explode_def by auto + +lemma nat_explode'_ne_suc: \\ n. nat_explode' (Suc n) \ nat_explode' n\ + by(rule exI [of _ \0::nat\], simp) + +lemma nat_explode'_digit: \hd (nat_explode' n ) < 10\ +proof(induct \n\) + case 0 + then show ?case by simp +next + case (Suc n) + then show ?case proof (cases \n < 9\) + case True + then show ?thesis by simp + next + case False + then show ?thesis + by simp + qed +qed + +lemma div_ten_less: \n \ 0 \ ((n::nat) div 10) < n\ + by simp + +lemma unroll_nat_explode': + \\ n < 10 \ (case n < 10 of True \ [n mod 10] | False \ n mod 10 # nat_explode' (n div 10)) = + (n mod 10 # nat_explode' (n div 10))\ + by simp + +lemma nat_explode_mod_10_ident: \map (\ x. x mod 10) (nat_explode' n) = nat_explode' n\ +proof (cases \n < 10\) + case True + then show ?thesis by simp +next + case False + then show ?thesis + proof (induct \n\ rule: nat_less_induct) + case (1 n) note * = this + then show ?case + using div_ten_less apply(simp (no_asm)) + using unroll_nat_explode'[of n] * + by (smt (z3) list.simps(8) list.simps(9) mod_div_trivial mod_eq_self_iff_div_eq_0 + nat_explode'.simps zero_less_numeral) + qed +qed + +lemma nat_explode'_digits: + \\d \ set (nat_explode' n). d < 10\ +proof + fix d + assume \d \ set (nat_explode' n)\ + then have \d \ set (map (\m. m mod 10) (nat_explode' n))\ + by (simp only: nat_explode_mod_10_ident) + then show \d < 10\ + by auto +qed + +lemma nat_explode_digits: + \\d \ set (nat_explode n). d < 10\ + using nat_explode'_digits [of n] by (simp only: nat_explode_def set_rev) + +value \nat_implode(nat_explode 42) = 42\ +value \nat_explode (Suc 21)\ + + +lemma nat_implode_append: + \nat_implode (a@[b]) = (1*b + foldr (+) (map (\(p, y). 10 ^ p * y) (enumerate (Suc 0) (rev a))) 0 )\ + by(simp add:nat_implode_def) + +lemma enumerate_suc: \enumerate (Suc n) l = map (\ (a,b). (a+1::nat,b)) (enumerate n l)\ +proof (induction \l\) + case Nil + then show ?case by simp +next + case (Cons a x) note * = this + then show ?case apply(simp only:enumerate_simps) + + apply(simp only:\enumerate (Suc n) x = map (\a. case a of (a, b) \ (a + 1, b)) (enumerate n x)\[symmetric]) + apply(simp) + using * + by (metis apfst_conv cond_case_prod_eta enumerate_Suc_eq) +qed + +lemma mult_assoc_aux1: + \(\(p, y). 10 ^ p * y) \ (\(a, y). (Suc a, y)) = (\(p, y). (10::nat) * (10 ^ p) * y)\ + by(auto simp:o_def) + +lemma fold_map_transfer: + \(foldr (+) (map (\(x,y). 10 * (f (x,y))) l) (0::nat)) = 10 * (foldr (+) (map (\x. (f x)) l) (0::nat))\ +proof(induct \l\) + case Nil + then show ?case by(simp) +next + case (Cons a l) + then show ?case by(simp) +qed + +lemma mult_assoc_aux2: \(\(p, y). 10 * 10 ^ p * (y::nat)) = (\(p, y). 10 * (10 ^ p * y))\ + by(auto) + +lemma nat_implode_explode_id: \nat_implode (nat_explode n) = n\ +proof (cases \n=0\) + case True note * = this + then show ?thesis + by (simp add: nat_explode_def nat_implode_def) +next + case False + then show ?thesis + proof (induct \n\ rule: nat_less_induct) + case (1 n) note * = this + then have + **: \nat_implode (nat_explode (n div 10)) = n div 10\ + proof (cases \n div 10 = 0\) + case True + then show ?thesis by(simp add: nat_explode_def nat_implode_def) + next + case False + then show ?thesis + using div_ten_less[of \n\] * + by(simp) + qed + then show ?case + proof (cases \n < 10\) + case True + then show ?thesis by(simp add: nat_explode_def nat_implode_def) + next + case False note *** = this + then show ?thesis + apply(simp (no_asm) add:nat_explode_def del:rev_rev_ident) + apply(simp only: bool.case(2)) + apply(simp del:nat_explode'.simps rev_rev_ident) + apply(fold nat_explode_def) + apply(simp only:nat_implode_append) + apply(simp add:enumerate_suc) + apply(simp only:mult_assoc_aux1) + using mult_assoc_aux2 apply(simp) + using fold_map_transfer[of \\(p, y). 10 ^ p * y\ + \(enumerate 0 (rev (nat_explode (n div 10))))\, simplified] + apply(simp) apply(fold nat_implode_def) using ** + by simp + qed + qed +qed + +definition + Read\<^sub>n\<^sub>a\<^sub>t :: \string \ nat\ where + \Read\<^sub>n\<^sub>a\<^sub>t s = nat_implode (map nat_of_digit s)\ + +definition + Show\<^sub>n\<^sub>a\<^sub>t::"nat \ string" where + \Show\<^sub>n\<^sub>a\<^sub>t n = map digit_of_nat (nat_explode n)\ + +declare Read\<^sub>n\<^sub>a\<^sub>t_def [solidity_symbex] + Show\<^sub>n\<^sub>a\<^sub>t_def [solidity_symbex] + +definition + \STR_is_nat s = (Show\<^sub>n\<^sub>a\<^sub>t (Read\<^sub>n\<^sub>a\<^sub>t s) = s)\ + +value \Read\<^sub>n\<^sub>a\<^sub>t ''10''\ +value \Show\<^sub>n\<^sub>a\<^sub>t 10\ +value \Read\<^sub>n\<^sub>a\<^sub>t (Show\<^sub>n\<^sub>a\<^sub>t (10)) = 10\ +value \Show\<^sub>n\<^sub>a\<^sub>t (Read\<^sub>n\<^sub>a\<^sub>t (''10'')) = ''10''\ + +lemma Show_nat_not_neg: + \set (Show\<^sub>n\<^sub>a\<^sub>t n) \{CHR ''0'', CHR ''1'', CHR ''2'', CHR ''3'', CHR ''4'', + CHR ''5'', CHR ''6'', CHR ''7'', CHR ''8'', CHR ''9''}\ + unfolding Show\<^sub>n\<^sub>a\<^sub>t_def + using nat_explode_digits[of n] img_digit_of_nat imageE image_set subsetI + by (smt (verit) imageE image_set subsetI) + +lemma Show_nat_not_empty: \(Show\<^sub>n\<^sub>a\<^sub>t n) \ []\ + by (simp add: Show\<^sub>n\<^sub>a\<^sub>t_def nat_explode_not_empty) + +lemma not_hd: \L \ [] \ e \ set(L) \ hd L \ e\ + by auto + +lemma Show_nat_not_neg'': \hd (Show\<^sub>n\<^sub>a\<^sub>t n) \ (CHR ''-'')\ + using Show_nat_not_neg[of \n\] + Show_nat_not_empty[of \n\] not_hd[of \Show\<^sub>n\<^sub>a\<^sub>t n\] + by auto + +lemma Show_Read_nat_id: \STR_is_nat s \ (Show\<^sub>n\<^sub>a\<^sub>t (Read\<^sub>n\<^sub>a\<^sub>t s) = s)\ + by(simp add:STR_is_nat_def) + +lemma bar': \\ d \ set l . d < 10 \ map nat_of_digit (map digit_of_nat l) = l\ + using nat_of_digit_digit_of_nat_id + by (simp add: map_idI) + +lemma Read_Show_nat_id: \Read\<^sub>n\<^sub>a\<^sub>t(Show\<^sub>n\<^sub>a\<^sub>t n) = n\ + apply(unfold Read\<^sub>n\<^sub>a\<^sub>t_def Show\<^sub>n\<^sub>a\<^sub>t_def) + using bar' nat_of_digit_digit_of_nat_id nat_explode_digits + using nat_implode_explode_id + by presburger + +definition + ReadL\<^sub>n\<^sub>a\<^sub>t :: \String.literal \ nat\ (\\_\\) where + \ReadL\<^sub>n\<^sub>a\<^sub>t = Read\<^sub>n\<^sub>a\<^sub>t \ String.explode\ + +definition + ShowL\<^sub>n\<^sub>a\<^sub>t::\nat \ String.literal\ (\\_\\)where + \ShowL\<^sub>n\<^sub>a\<^sub>t = String.implode \ Show\<^sub>n\<^sub>a\<^sub>t\ + +declare ReadL\<^sub>n\<^sub>a\<^sub>t_def [solidity_symbex] + ShowL\<^sub>n\<^sub>a\<^sub>t_def [solidity_symbex] + + +definition + \strL_is_nat' s = (ShowL\<^sub>n\<^sub>a\<^sub>t (ReadL\<^sub>n\<^sub>a\<^sub>t s) = s)\ + +value \\STR ''10''\::nat\ +value \ReadL\<^sub>n\<^sub>a\<^sub>t (STR ''10'')\ +value \\10::nat\\ +value \ShowL\<^sub>n\<^sub>a\<^sub>t 10\ +value \ReadL\<^sub>n\<^sub>a\<^sub>t (ShowL\<^sub>n\<^sub>a\<^sub>t (10)) = 10\ +value \ShowL\<^sub>n\<^sub>a\<^sub>t (ReadL\<^sub>n\<^sub>a\<^sub>t (STR ''10'')) = STR ''10''\ + +lemma Show_Read_nat'_id: \strL_is_nat' s \ (ShowL\<^sub>n\<^sub>a\<^sub>t (ReadL\<^sub>n\<^sub>a\<^sub>t s) = s)\ + by(simp add:strL_is_nat'_def) + +lemma digits_are_ascii: + \c \ {CHR ''0'', CHR ''1'', CHR ''2'', CHR ''3'', CHR ''4'', + CHR ''5'', CHR ''6'', CHR ''7'', CHR ''8'', CHR ''9''} + \ String.ascii_of c = c\ + by auto + +lemma Show\<^sub>n\<^sub>a\<^sub>t_ascii: \map String.ascii_of (Show\<^sub>n\<^sub>a\<^sub>t n) = Show\<^sub>n\<^sub>a\<^sub>t n\ + using Show_nat_not_neg digits_are_ascii + by (meson map_idI subsetD) + + +lemma Read_Show_nat'_id: \ReadL\<^sub>n\<^sub>a\<^sub>t(ShowL\<^sub>n\<^sub>a\<^sub>t n) = n\ + apply(unfold ReadL\<^sub>n\<^sub>a\<^sub>t_def ShowL\<^sub>n\<^sub>a\<^sub>t_def, simp) + by (simp add: Show\<^sub>n\<^sub>a\<^sub>t_ascii Read_Show_nat_id) + + +subsubsection\Integer\ +definition + Read\<^sub>i\<^sub>n\<^sub>t :: \string \ int\ where + \Read\<^sub>i\<^sub>n\<^sub>t x = (if hd x = (CHR ''-'') then -(int (Read\<^sub>n\<^sub>a\<^sub>t (tl x))) else int (Read\<^sub>n\<^sub>a\<^sub>t x))\ + +definition + Show\<^sub>i\<^sub>n\<^sub>t::\int \ string\ where + \Show\<^sub>i\<^sub>n\<^sub>t i = (if i < 0 then (CHR ''-'')#(Show\<^sub>n\<^sub>a\<^sub>t (nat (-i))) + else Show\<^sub>n\<^sub>a\<^sub>t (nat i))\ + +definition + \STR_is_int s = (Show\<^sub>i\<^sub>n\<^sub>t (Read\<^sub>i\<^sub>n\<^sub>t s) = s)\ + + +declare Read\<^sub>i\<^sub>n\<^sub>t_def [solidity_symbex] + Show\<^sub>i\<^sub>n\<^sub>t_def [solidity_symbex] + +value \Read\<^sub>i\<^sub>n\<^sub>t (Show\<^sub>i\<^sub>n\<^sub>t 10) = 10\ +value \Read\<^sub>i\<^sub>n\<^sub>t (Show\<^sub>i\<^sub>n\<^sub>t (-10)) = -10\ + +value \Show\<^sub>i\<^sub>n\<^sub>t (Read\<^sub>i\<^sub>n\<^sub>t (''10'')) = ''10''\ +value \Show\<^sub>i\<^sub>n\<^sub>t (Read\<^sub>i\<^sub>n\<^sub>t (''-10'')) = ''-10''\ + +lemma Show_Read_id: \STR_is_int s \ (Show\<^sub>i\<^sub>n\<^sub>t (Read\<^sub>i\<^sub>n\<^sub>t s) = s)\ + by(simp add:STR_is_int_def) + +lemma Read_Show_id: \Read\<^sub>i\<^sub>n\<^sub>t(Show\<^sub>i\<^sub>n\<^sub>t(x)) = x\ + apply(code_simp) + apply(auto simp:Show_nat_not_neg Read_Show_nat_id)[1] + apply(thin_tac \\ x < 0 \) + using Show_nat_not_neg'' + by blast + +lemma STR_is_int_Show: \STR_is_int (Show\<^sub>i\<^sub>n\<^sub>t n)\ + by(simp add:STR_is_int_def Read_Show_id) + +definition + ReadL\<^sub>i\<^sub>n\<^sub>t :: \String.literal \ int\ (\\_\\) where + \ReadL\<^sub>i\<^sub>n\<^sub>t = Read\<^sub>i\<^sub>n\<^sub>t \ String.explode\ + +definition + ShowL\<^sub>i\<^sub>n\<^sub>t::\int \ String.literal\ (\\_\\) where + \ShowL\<^sub>i\<^sub>n\<^sub>t =String.implode \ Show\<^sub>i\<^sub>n\<^sub>t\ + +definition + \strL_is_int' s = (ShowL\<^sub>i\<^sub>n\<^sub>t (ReadL\<^sub>i\<^sub>n\<^sub>t s) = s)\ + +declare ReadL\<^sub>i\<^sub>n\<^sub>t_def [solidity_symbex] + ShowL\<^sub>i\<^sub>n\<^sub>t_def [solidity_symbex] + +value \ReadL\<^sub>i\<^sub>n\<^sub>t (ShowL\<^sub>i\<^sub>n\<^sub>t 10) = 10\ +value \ReadL\<^sub>i\<^sub>n\<^sub>t (ShowL\<^sub>i\<^sub>n\<^sub>t (-10)) = -10\ + +value \ShowL\<^sub>i\<^sub>n\<^sub>t (ReadL\<^sub>i\<^sub>n\<^sub>t (STR ''10'')) = STR ''10''\ +value \ShowL\<^sub>i\<^sub>n\<^sub>t (ReadL\<^sub>i\<^sub>n\<^sub>t (STR ''-10'')) = STR ''-10''\ + +lemma Show_ReadL_id: \strL_is_int' s \ (ShowL\<^sub>i\<^sub>n\<^sub>t (ReadL\<^sub>i\<^sub>n\<^sub>t s) = s)\ + by(simp add:strL_is_int'_def) + +lemma Read_ShowL_id: \ReadL\<^sub>i\<^sub>n\<^sub>t (ShowL\<^sub>i\<^sub>n\<^sub>t x) = x\ +proof(cases \x < 0\) + case True + then show ?thesis using ShowL\<^sub>i\<^sub>n\<^sub>t_def ReadL\<^sub>i\<^sub>n\<^sub>t_def Show\<^sub>i\<^sub>n\<^sub>t_def Show\<^sub>n\<^sub>a\<^sub>t_ascii + by (metis (no_types, lifting) Read_Show_id String.ascii_of_Char comp_def implode.rep_eq list.simps(9)) +next + case False + then show ?thesis using ShowL\<^sub>i\<^sub>n\<^sub>t_def ReadL\<^sub>i\<^sub>n\<^sub>t_def Show\<^sub>i\<^sub>n\<^sub>t_def Show\<^sub>n\<^sub>a\<^sub>t_ascii + by (metis Read_Show_id String.explode_implode_eq comp_apply) +qed + +lemma STR_is_int_ShowL: \strL_is_int' (ShowL\<^sub>i\<^sub>n\<^sub>t n)\ + by(simp add:strL_is_int'_def Read_ShowL_id) + +lemma String_Cancel: "a + (c::String.literal) = b + c \ a = b" +using String.plus_literal.rep_eq +by (metis append_same_eq literal.explode_inject) + +end + diff --git a/thys/Solidity/Reentrancy.thy b/thys/Solidity/Reentrancy.thy --- a/thys/Solidity/Reentrancy.thy +++ b/thys/Solidity/Reentrancy.thy @@ -1,3308 +1,840 @@ -section\Reentrancy\ - -text \ - In the following we use our semantics to verify a contract implementing a simple token. - The contract is defined by definition \<^emph>\victim\ and consist of one state variable and two methods: - \<^item> The state variable "balance" is a mapping which assigns a balance to each address. - \<^item> Method "deposit" allows to send money to the contract which is then added to the sender's balance. - \<^item> Method "withdraw" allows to withdraw the callers balance. -\ - -text \ - We then verify that the following invariant (defined by \<^emph>\INV\) is preserved by both methods: - The difference between - \<^item> the contracts own account-balance and - \<^item> the sum of all the balances kept in the contracts state variable - is larger than a certain threshold. -\ - -text \ - There are two things to note here: - First, Solidity implicitly triggers the call of a so-called fallback method whenever we transfer money to a contract. - In particular if another contract calls "withdraw", this triggers an implict call to the callee's fallback method. - This functionality was exploited in the infamous DAO attack which we demonstrate it in terms of an example later on. - Since we do not know all potential contracts which call "withdraw", we need to verify our invariant for all possible Solidity programs. - Thus, the core result here is a lemma which shows that the invariant is preserved by every Solidity program which is not executed in the context of our own contract. - For our own methods we show that the invariant holds after executing it. - Since our own program as well as the unknown program may depend on each other both properties are encoded in a single lemma (\<^emph>\secure\) which is then proved by induction over all statements. - The final result is then given in terms of two corollaries for the corresponding methods of our contract. -\ - -text \ - The second thing to note is that we were not able to verify that the difference is indeed constant. - During verification it turned out that this is not the case since in the fallback method a contract could just send us additional money withouth calling "deposit". - In such a case the difference would change. In particular it would grow. However, we were able to verify that the difference does never shrink which is what we actually want to ensure. -\ - -theory Reentrancy -imports Solidity_Evaluator -begin - -subsection\Example of Re-entrancy\ - -(* The following value can take several minutes to process. *) -value "eval 1000 - stmt - (COMP - (EXTERNAL (ADDRESS (STR ''Victim'')) (STR ''deposit'') [] (UINT 256 10)) - (EXTERNAL (ADDRESS (STR ''Victim'')) (STR ''withdraw'') [] (UINT 256 0))) - (STR ''Attacker'') - (STR '''') - (STR ''0'') - [(STR ''Victim'', STR ''100''), (STR ''Attacker'', STR ''100'')] - [ - (STR ''Attacker'', - [], - ITE - (LESS (BALANCE THIS) (UINT 256 125)) - (EXTERNAL (ADDRESS (STR ''Victim'')) (STR ''withdraw'') [] (UINT 256 0)) - SKIP), - (STR ''Victim'', - [ - (STR ''balance'', Var (STMap TAddr (STValue (TUInt 256)))), - (STR ''deposit'', Method ([], ASSIGN (Ref (STR ''balance'') [SENDER]) VALUE, None)), - (STR ''withdraw'', Method ([], - ITE - (LESS (UINT 256 0) (LVAL (Ref (STR ''balance'') [SENDER]))) - (COMP - (TRANSFER SENDER (LVAL (Ref (STR ''balance'') [SENDER]))) - (ASSIGN (Ref (STR ''balance'') [SENDER]) (UINT 256 0))) - SKIP - , None))], - SKIP) - ] - []" - -subsection\Definition of Contract\ - -abbreviation myrexp::L - where "myrexp \ Ref (STR ''balance'') [SENDER]" - -abbreviation mylval::E - where "mylval \ LVAL myrexp" - -abbreviation assign::S - where "assign \ ASSIGN (Ref (STR ''balance'') [SENDER]) (UINT 256 0)" - -abbreviation transfer::S - where "transfer \ TRANSFER SENDER (LVAL (Id (STR ''bal'')))" - -abbreviation comp::S - where "comp \ COMP assign transfer" - -abbreviation keep::S - where "keep \ BLOCK ((STR ''bal'', Value (TUInt 256)), Some mylval) comp" - -abbreviation deposit::S - where "deposit \ ASSIGN (Ref (STR ''balance'') [SENDER]) (PLUS (LVAL (Ref (STR ''balance'') [SENDER])) VALUE)" - -definition victim::"(Identifier, Member) fmap" - where "victim \ fmap_of_list [ - (STR ''balance'', Var (STMap TAddr (STValue (TUInt 256)))), - (STR ''deposit'', Method ([], deposit, None)), - (STR ''withdraw'', Method ([], keep, None))]" - -subsection\Definition of Invariant\ - -abbreviation "SUMM s \ \(ad,x)|fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x. ReadL\<^sub>i\<^sub>n\<^sub>t x" - -abbreviation "POS s \ \ad x. fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x \ ReadL\<^sub>i\<^sub>n\<^sub>t x \ 0" - -abbreviation "INV st s val bal \ - fmlookup (storage st) (STR ''Victim'') = Some s \ ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st) (STR ''Victim'')) - val \ bal \ bal \ 0" - -definition frame_def: "frame bal st \ (\s. INV st s (SUMM s) bal \ POS s)" - -subsection\Verification\ - -lemma conj3: "P \ Q \ R \ P \ (Q \ R)" by simp - -lemma fmfinite: "finite ({(ad, x). fmlookup y ad = Some x})" -proof - - have "{(ad, x). fmlookup y ad = Some x} \ dom (fmlookup y) \ ran (fmlookup y)" - proof - fix x assume "x \ {(ad, x). fmlookup y ad = Some x}" - then have "fmlookup y (fst x) = Some (snd x)" by auto - then have "fst x \ dom (fmlookup y)" and "snd x \ ran (fmlookup y)" using Map.ranI by (blast,metis) - then show "x \ dom (fmlookup y) \ ran (fmlookup y)" by (simp add: mem_Times_iff) - qed - thus ?thesis by (simp add: finite_ran finite_subset) -qed - -lemma fmlookup_finite: - fixes f :: "'a \ 'a" - and y :: "('a, 'b) fmap" - assumes "inj_on (\(ad, x). (f ad, x)) {(ad, x). (fmlookup y \ f) ad = Some x}" - shows "finite {(ad, x). (fmlookup y \ f) ad = Some x}" -proof (cases rule: inj_on_finite[OF assms(1), of "{(ad, x)|ad x. (fmlookup y) ad = Some x}"]) - case 1 - then show ?case by auto -next - case 2 - then have *: "finite {(ad, x) |ad x. fmlookup y ad = Some x}" using fmfinite[of y] by simp - show ?case using finite_image_set[OF *, of "\(ad, x). (ad, x)"] by simp -qed - -lemma balance_inj: "inj_on (\(ad, x). (ad + (STR ''.'' + STR ''balance''),x)) {(ad, x). (fmlookup y \ f) ad = Some x}" -proof - fix v v' assume asm1: "v \ {(ad, x). (fmlookup y \ f) ad = Some x}" and asm2: "v' \ {(ad, x). (fmlookup y \ f) ad = Some x}" - and *:"(case v of (ad, x) \ (ad + (STR ''.'' + STR ''balance''),x)) = (case v' of (ad, x) \ (ad + (STR ''.'' + STR ''balance''),x))" - then obtain ad ad' x x' where **: "v = (ad,x)" and ***: "v'=(ad',x')" by auto - moreover from * ** *** have "ad + (STR ''.'' + STR ''balance'') = ad' + (STR ''.'' + STR ''balance'')" by simp - with * ** have "ad = ad'" using String_Cancel[of ad "STR ''.'' + STR ''balance''" ad' ] by simp - moreover from asm1 asm2 ** *** have "(fmlookup y \ f) ad = Some x" and "(fmlookup y \ f) ad' = Some x'" by auto - with calculation(3) have "x=x'" by simp - ultimately show "v=v'" by simp -qed - -lemma transfer_frame: -assumes "Accounts.transfer ad adv v (accounts st) = Some acc" - and "frame bal st" - and "ad \ STR ''Victim''" - shows "frame bal (st\accounts := acc\)" -proof (cases "adv = STR ''Victim''") - case True - define st' where "st' = st\accounts := acc, stack := emptyStore, memory := emptyStore\" - from True assms(2) transfer_mono[OF assms(1)] have "(\s. fmlookup (storage st) (STR ''Victim'') = Some s \ ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance acc (STR ''Victim'')) - (SUMM s) \ bal \ bal \ 0)" by (auto simp add:frame_def) - then have "(\s. fmlookup (storage st') (STR ''Victim'') = Some s \ ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st') (STR ''Victim'')) - (SUMM s) \ bal \ bal \ 0)" by (auto simp add: st'_def) - then show ?thesis using assms(2) by (auto simp add:st'_def frame_def) -next - case False - define st' where "st' = st\accounts := acc, stack := emptyStore, memory := emptyStore\" - from False assms(2) assms(3) transfer_eq[OF assms(1)] have "(\s. fmlookup (storage st) (STR ''Victim'') = Some s \ ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance acc (STR ''Victim'')) - (SUMM s) \ bal \ bal \ 0)" by (auto simp add:frame_def) - then have "(\s. fmlookup (storage st') (STR ''Victim'') = Some s \ ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st') (STR ''Victim'')) - (SUMM s) \ bal \ bal \ 0)" by (auto simp add: st'_def) - then show ?thesis using assms(2) by (auto simp add:st'_def frame_def) -qed - -lemma decl_frame: - assumes "frame bal st" - and "decl a1 a2 a3 cp cd mem c env st = Normal (rv, st')" - shows "frame bal st'" -proof (cases a2) - case (Value t) - then show ?thesis - proof (cases a3) - case None - with Value show ?thesis using assms by (auto simp add:frame_def) - next - case (Some a) - show ?thesis - proof (cases a) - case (Pair a b) - then show ?thesis - proof (cases a) - case (KValue v) - then show ?thesis - proof (cases b) - case v2: (Value t') - show ?thesis - proof (cases "Valuetypes.convert t' t v") - case None - with Some Pair KValue Value v2 show ?thesis using assms by simp - next - case s2: (Some a) - show ?thesis - proof (cases a) - case p2: (Pair a b) - with Some Pair KValue Value v2 s2 show ?thesis using assms by (auto simp add:frame_def) - qed - qed - next - case (Calldata x2) - with Some Pair KValue Value show ?thesis using assms by simp - next - case (Memory x3) - with Some Pair KValue Value show ?thesis using assms by simp - next - case (Storage x4) - with Some Pair KValue Value show ?thesis using assms by simp - qed - next - case (KCDptr x2) - with Some Pair Value show ?thesis using assms by simp - next - case (KMemptr x3) - with Some Pair Value show ?thesis using assms by simp - next - case (KStoptr x4) - with Some Pair Value show ?thesis using assms by simp - qed - qed - qed -next - case (Calldata x2) - then show ?thesis - proof (cases cp) - case True - then show ?thesis - proof (cases x2) - case (MTArray x t) - then show ?thesis - proof (cases a3) - case None - with Calldata show ?thesis using assms by simp - next - case (Some a) - show ?thesis - proof (cases a) - case (Pair a b) - then show ?thesis - proof (cases a) - case (KValue x1) - with Calldata Some Pair show ?thesis using assms by simp - next - case (KCDptr p) - define l where "l = ShowL\<^sub>n\<^sub>a\<^sub>t (toploc c)" - obtain c' where c_def: "\x. (x, c') = allocate c" by simp - show ?thesis - proof (cases "cpm2m p l x t cd c'") - case None - with Calldata MTArray Some Pair KCDptr l_def c_def True show ?thesis using assms by simp - next - case s2: (Some a) - with Calldata MTArray Some Pair KCDptr l_def c_def True show ?thesis using assms by (auto simp add:frame_def) - qed - next - case (KMemptr p) - define l where "l = ShowL\<^sub>n\<^sub>a\<^sub>t (toploc c)" - obtain c' where c_def: "\x. (x, c') = allocate c" by simp - show ?thesis - proof (cases "cpm2m p l x t mem c'") - case None - with Calldata MTArray Some Pair KMemptr l_def c_def True show ?thesis using assms by simp - next - case s2: (Some a) - with Calldata MTArray Some Pair KMemptr l_def c_def True show ?thesis using assms by (auto simp add:frame_def) - qed - next - case (KStoptr x4) - with Calldata Some Pair show ?thesis using assms by simp - qed - qed - qed - next - case (MTValue x2) - with Calldata show ?thesis using assms by simp - qed - next - case False - with Calldata show ?thesis using assms by simp - qed -next - case (Memory x3) - then show ?thesis - proof (cases x3) - case (MTArray x t) - then show ?thesis - proof (cases a3) - case None - with Memory MTArray None show ?thesis using assms by (auto simp add:frame_def simp add:Let_def) - next - case (Some a) - then show ?thesis - proof (cases a) - case (Pair a b) - then show ?thesis - proof (cases a) - case (KValue x1) - with Memory MTArray Some Pair show ?thesis using assms by simp - next - case (KCDptr p) - define m l where "m = memory st" and "l = ShowL\<^sub>n\<^sub>a\<^sub>t (toploc m)" - obtain m' where m'_def: "\x. (x, m') = allocate m" by simp - then show ?thesis - proof (cases "cpm2m p l x t cd m'") - case None - with Memory MTArray Some Pair KCDptr m_def l_def m'_def show ?thesis using assms by simp - next - case s2: (Some a) - with Memory MTArray Some Pair KCDptr m_def l_def m'_def show ?thesis using assms by (auto simp add:frame_def) - qed - next - case (KMemptr p) - then show ?thesis - proof (cases cp) - case True - define m l where "m = memory st" and "l = ShowL\<^sub>n\<^sub>a\<^sub>t (toploc m)" - obtain m' where m'_def: "\x. (x, m') = allocate m" by simp - then show ?thesis - proof (cases "cpm2m p l x t mem m'") - case None - with Memory MTArray Some Pair KMemptr True m_def l_def m'_def show ?thesis using assms by simp - next - case s2: (Some a) - with Memory MTArray Some Pair KMemptr True m_def l_def m'_def show ?thesis using assms by (auto simp add:frame_def) - qed - next - case False - with Memory MTArray Some Pair KMemptr show ?thesis using assms by (auto simp add:frame_def) - qed - next - case (KStoptr p) - then show ?thesis - proof (cases b) - case (Value x1) - with Memory MTArray Some Pair KStoptr show ?thesis using assms by simp - next - case (Calldata x2) - with Memory MTArray Some Pair KStoptr show ?thesis using assms by simp - next - case m2: (Memory x3) - with Memory MTArray Some Pair KStoptr show ?thesis using assms by simp - next - case (Storage x4) - then show ?thesis - proof (cases x4) - case (STArray x' t') - define m l where "m = memory st" and "l = ShowL\<^sub>n\<^sub>a\<^sub>t (toploc m)" - obtain m' where m'_def: "\x. (x, m') = allocate m" by simp - from assms(2) Memory MTArray Some Pair KStoptr Storage STArray m_def l_def m'_def - obtain s where *: "fmlookup (storage st) (address env) = Some s" using Let_def by (auto simp add: Let_def split:option.split_asm) - then show ?thesis - proof (cases "cps2m p l x' t' s m'") - case None - with Memory MTArray Some Pair KStoptr Storage STArray m_def l_def m'_def * show ?thesis using assms by simp - next - case s2: (Some a) - with Memory MTArray Some Pair KStoptr Storage STArray m_def l_def m'_def * show ?thesis using assms by (auto simp add:frame_def) - qed - next - case (STMap x21 x22) - with Memory MTArray Some Pair KStoptr Storage show ?thesis using assms by simp - next - case (STValue x3) - with Memory MTArray Some Pair KStoptr Storage show ?thesis using assms by simp - qed - qed - qed - qed - qed - next - case (MTValue x2) - with Memory show ?thesis using assms by simp - qed -next - case (Storage x4) - then show ?thesis - proof (cases x4) - case (STArray x t) - then show ?thesis - proof (cases a3) - case None - with Storage STArray show ?thesis using assms by simp - next - case (Some a) - then show ?thesis - proof (cases a) - case (Pair a b) - then show ?thesis - proof (cases a) - case (KValue x1) - with Storage STArray Some Pair show ?thesis using assms by simp - next - case (KCDptr x2) - with Storage STArray Some Pair show ?thesis using assms by simp - next - case (KMemptr x3) - with Storage STArray Some Pair show ?thesis using assms by simp - next - case (KStoptr x4) - with Storage STArray Some Pair show ?thesis using assms by (auto simp add:frame_def) - qed - qed - qed - next - case (STMap t t') - then show ?thesis - proof (cases a3) - case None - with Storage STMap show ?thesis using assms by simp - next - case (Some a) - then show ?thesis - proof (cases a) - case (Pair a b) - then show ?thesis - proof (cases a) - case (KValue x1) - with Storage STMap Some Pair show ?thesis using assms by simp - next - case (KCDptr x2) - with Storage STMap Some Pair show ?thesis using assms by simp - next - case (KMemptr x3) - with Storage STMap Some Pair show ?thesis using assms by simp - next - case (KStoptr x4) - with Storage STMap Some Pair show ?thesis using assms by (auto simp add:frame_def) - qed - qed - qed - next - case (STValue x3) - with Storage show ?thesis using assms by simp - qed -qed - - -context statement_with_gas -begin - -lemma secureassign: - assumes "stmt assign ep env cd st = Normal((), st')" - and "fmlookup (storage st) (STR ''Victim'') = Some s" - and "address env = (STR ''Victim'')" - and "fmlookup (denvalue env) (STR ''balance'') = Some (Storage (STMap TAddr (STValue (TUInt 256))), Storeloc STR ''balance'')" - and "accessStore x (stack st) = Some (KValue (accessStorage (TUInt 256) (sender env + (STR ''.'' + STR ''balance'')) s))" - and "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st) (STR ''Victim'')) - (SUMM s) \ bal" - and "POS s" - obtains s' - where "fmlookup (storage st') (STR ''Victim'') = Some s'" - and "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st') (STR ''Victim'')) - (SUMM s' + ReadL\<^sub>i\<^sub>n\<^sub>t (accessStorage (TUInt 256) (sender env + (STR ''.'' + STR ''balance'')) s)) \ bal" - and "accessStore x (stack st') = Some (KValue (accessStorage (TUInt 256) (sender env + (STR ''.'' + STR ''balance'')) s))" - and "POS s'" -proof - - define st'' where "st'' = st\gas := gas st - costs assign ep env cd st\" - define st''' where "st''' = st''\gas := gas st'' - costs\<^sub>e (UINT 256 0) ep env cd st''\" - define st'''' where "st'''' = st'''\gas := gas st''' - costs\<^sub>e SENDER ep env cd st'''\" - - from assms(1) have c1: "gas st > costs assign ep env cd st" by (auto split:if_split_asm) - have c2: "gas st'' > costs\<^sub>e (UINT 256 0) ep env cd st''" - proof (rule ccontr) - assume "\ costs\<^sub>e (UINT 256 0) ep env cd st'' < gas st''" - with c1 show False using assms(1) st''_def st'''_def by auto - qed - hence *:"expr (UINT 256 0) ep env cd st'' = Normal ((KValue (createUInt 256 0),Value (TUInt 256)), st''')" using expr.simps(2)[of 256 0 ep env cd "st\gas := gas st - costs assign ep env cd st\"] st''_def st'''_def by simp - moreover have "gas st''' > costs\<^sub>e SENDER ep env cd st'''" - proof (rule ccontr) - assume "\ costs\<^sub>e SENDER ep env cd st''' < gas st'''" - with c1 c2 show False using assms(1,4) st''_def st'''_def by auto - qed - with assms(4) have **:"lexp (Ref (STR ''balance'') [SENDER]) ep env cd st''' = Normal ((LStoreloc ((sender env) + (STR ''.'' + STR ''balance'')), Storage (STValue (TUInt 256))), st'''')" using st''''_def by simp - moreover have "Valuetypes.convert (TUInt 256) (TUInt 256) (ShowL\<^sub>i\<^sub>n\<^sub>t 0) = Some (ShowL\<^sub>i\<^sub>n\<^sub>t 0, TUInt 256)" by simp - moreover from * ** st''_def assms(1) obtain s'' where ***: "fmlookup (storage st'''') (address env) = Some s''" by (auto split:if_split_asm option.split_asm) - ultimately have ****:"st' = st''''\storage := fmupd (STR ''Victim'') (fmupd ((sender env) + (STR ''.'' + STR ''balance'')) (ShowL\<^sub>i\<^sub>n\<^sub>t 0) s'') (storage st)\" using c1 st''_def st'''_def st''''_def assms(1,3) by auto - moreover define s' where s'_def: "s' = (fmupd ((sender env) + (STR ''.'' + STR ''balance'')) (ShowL\<^sub>i\<^sub>n\<^sub>t 0) s'')" - ultimately have "fmlookup (storage st') (STR ''Victim'') = Some s'" - and *****:"fmlookup s' ((sender env) + (STR ''.'' + STR ''balance'')) = Some (ShowL\<^sub>i\<^sub>n\<^sub>t 0)" by simp_all - moreover have "SUMM s' + ReadL\<^sub>i\<^sub>n\<^sub>t (accessStorage (TUInt 256) (sender env + (STR ''.'' + STR ''balance'')) s) = SUMM s" - proof - - have s1: "SUMM s = (\(ad,x)|fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env. ReadL\<^sub>i\<^sub>n\<^sub>t x) + ReadL\<^sub>i\<^sub>n\<^sub>t (accessStorage (TUInt 256) (sender env + (STR ''.'' + STR ''balance'')) s)" - proof (cases "fmlookup s (sender env + (STR ''.'' + STR ''balance'')) = None") - case True - then have "accessStorage (TUInt 256) (sender env + (STR ''.'' + STR ''balance'')) s = ShowL\<^sub>i\<^sub>n\<^sub>t 0" by simp - moreover have "{(ad,x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x} = {(ad,x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env}" - proof - show "{(ad, x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x} \ {(ad, x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env}" - proof - fix x - assume "x \ {(ad, x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x}" - then show "x \ {(ad, x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env}" using True by auto - qed - next - show "{(ad, x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env} \ {(ad, x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x }" - proof - fix x - assume "x \ {(ad, x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env}" - then show "x \ {(ad, x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x}" using True by auto - qed - qed - then have "SUMM s = (\(ad,x)|fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env. ReadL\<^sub>i\<^sub>n\<^sub>t x)" by simp - ultimately show ?thesis using Read_ShowL_id by simp - next - case False - then obtain val where val_def: "fmlookup s (sender env + (STR ''.'' + STR ''balance'')) = Some val" by auto - - have "inj_on (\(ad, x). (ad + (STR ''.'' + STR ''balance''), x)) {(ad, x). (fmlookup s \ (\ad. ad + (STR ''.'' + STR ''balance''))) ad = Some x}" using balance_inj by simp - then have "finite {(ad, x). (fmlookup s \ (\ad. ad + (STR ''.'' + STR ''balance''))) ad = Some x}" using fmlookup_finite[of "\ad. (ad + (STR ''.'' + STR ''balance''))" s] by simp - then have sum1: "finite ({(ad,x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env})" using finite_subset[of "{(ad, x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env}" "{(ad, x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x}"] by auto - moreover have sum2: "(sender env,val) \ {(ad,x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env}" by simp - moreover from sum1 x1 val_def have "insert (sender env,val) {(ad, x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env} = {(ad, x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x}" by auto - ultimately show ?thesis using sum.insert[OF sum1 sum2, of "\(ad,x). ReadL\<^sub>i\<^sub>n\<^sub>t x"] val_def by simp - qed - moreover have s2: "SUMM s' = (\(ad,x)|fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env. ReadL\<^sub>i\<^sub>n\<^sub>t x)" - proof - - have "inj_on (\(ad, x). (ad + (STR ''.'' + STR ''balance''), x)) {(ad, x). (fmlookup s' \ (\ad. ad + (STR ''.'' + STR ''balance''))) ad = Some x}" using balance_inj by simp - then have "finite {(ad, x). (fmlookup s' \ (\ad. ad + (STR ''.'' + STR ''balance''))) ad = Some x}" using fmlookup_finite[of "\ad. (ad + (STR ''.'' + STR ''balance''))" s'] by simp - then have sum1: "finite ({(ad,x). fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env})" using finite_subset[of "{(ad, x). fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env}" "{(ad, x). fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x}"] by auto - moreover have sum2: "(sender env,ShowL\<^sub>i\<^sub>n\<^sub>t 0) \ {(ad,x). fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env}" by simp - moreover from ***** have "insert (sender env,ShowL\<^sub>i\<^sub>n\<^sub>t 0) {(ad, x). fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env} = {(ad, x). fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x}" by auto - ultimately show ?thesis using sum.insert[OF sum1 sum2, of "\(ad,x). ReadL\<^sub>i\<^sub>n\<^sub>t x"] Read_ShowL_id by simp - qed - moreover have s3: "(\(ad,x)|fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env. ReadL\<^sub>i\<^sub>n\<^sub>t x) - =(\(ad,x)|fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env. ReadL\<^sub>i\<^sub>n\<^sub>t x)" - proof - - have "{(ad,x). fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env} = {(ad,x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env}" - proof - show "{(ad, x). fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env} - \ {(ad, x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env}" - proof - fix xx - assume "xx \ {(ad, x). fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env}" - then obtain ad x where "xx = (ad,x)" and "fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x" and "ad \ sender env" by auto - moreover have "s''=s" using assms(2,3) s'_def *** st''''_def st'''_def st''_def by simp - moreover from `ad \ sender env` have "ad + (STR ''.'' + STR ''balance'') \ (sender env) + (STR ''.'' + STR ''balance'')" using String_Cancel[where c="(STR ''.'' + STR ''balance'')"] by auto - ultimately show "xx \ {(ad, x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env}" using s'_def by simp - qed - next - show "{(ad, x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env} - \ {(ad, x). fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env}" - proof - fix xx - assume "xx \ {(ad, x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env}" - then obtain ad x where "xx = (ad,x)" and "fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x" and "ad \ sender env" by auto - moreover have "s''=s" using assms(2,3) s'_def *** st''''_def st'''_def st''_def by simp - moreover from `ad \ sender env` have "ad + (STR ''.'' + STR ''balance'') \ (sender env) + (STR ''.'' + STR ''balance'')" using String_Cancel[where c="(STR ''.'' + STR ''balance'')"] by auto - ultimately show "xx \ {(ad, x). fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env}" using s'_def by simp - qed - qed - thus ?thesis by simp - qed - ultimately have "SUMM s' = SUMM s - ReadL\<^sub>i\<^sub>n\<^sub>t (accessStorage (TUInt 256) (sender env + (STR ''.'' + STR ''balance'')) s) " - proof - - from s2 have "SUMM s' = (\(ad,x)|fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env. ReadL\<^sub>i\<^sub>n\<^sub>t x)" by simp - also from s3 have "\ = (\(ad,x)|fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env. ReadL\<^sub>i\<^sub>n\<^sub>t x)" by simp - also from s1 have "\ = SUMM s - ReadL\<^sub>i\<^sub>n\<^sub>t (accessStorage (TUInt 256) (sender env + (STR ''.'' + STR ''balance'')) s) " by simp - finally show ?thesis . - qed - then show ?thesis by simp - qed - moreover have "POS s'" - proof (rule allI[OF allI]) - fix x xa - show "fmlookup s' (x + (STR ''.'' + STR ''balance'')) = Some xa \ 0 \ (\xa\::int)" - proof - assume a1: "fmlookup s' (x + (STR ''.'' + STR ''balance'')) = Some xa" - show "0 \ (\xa\::int)" - proof (cases "x = sender env") - case True - then show ?thesis using s'_def a1 using Read_ShowL_id by auto - next - case False - moreover have "s''=s" using assms(2,3) s'_def *** st''''_def st'''_def st''_def by simp - ultimately have "fmlookup s (x + (STR ''.'' + STR ''balance'')) = Some xa" using s'_def a1 String_Cancel by (auto split:if_split_asm) - then show ?thesis using assms(7) by simp - qed - qed - qed - moreover have "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st') (STR ''Victim'')) = ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st) (STR ''Victim''))" using **** st''_def st'''_def st''''_def by simp - moreover from assms(5) have "accessStore x (stack st') = Some (KValue (accessStorage (TUInt 256) (sender env + (STR ''.'' + STR ''balance'')) s) )"using **** st''_def st'''_def st''''_def by simp - ultimately show ?thesis using assms(6) that by simp -qed - - -lemma securesender: - assumes "expr SENDER ep env cd st = Normal((KValue v,t), st')" - and "fmlookup (storage st) (STR ''Victim'') = Some s" - and "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st) (STR ''Victim'')) - SUMM s \ bal \ POS s" - obtains s' where - "v = sender env" - and "t = Value TAddr" - and "fmlookup (storage st') (STR ''Victim'') = Some s'" - and "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st') (STR ''Victim'')) - SUMM s' \ bal \ POS s'" - using assms by (auto split:if_split_asm) - -lemma securessel: - assumes "ssel type loc [] ep env cd st = Normal (x, st')" - and "fmlookup (storage st) (STR ''Victim'') = Some s" - and "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st) (STR ''Victim'')) - SUMM s \ bal \ POS s" - obtains s' where - "x = (loc, type)" - and "fmlookup (storage st') (STR ''Victim'') = Some s'" - and "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st') (STR ''Victim'')) - SUMM s' \ bal \ POS s'" - using assms by auto - -lemma securessel2: - assumes "ssel (STMap TAddr (STValue (TUInt 256))) (STR ''balance'') [SENDER] ep env cd st = Normal ((loc, type), st')" - and "fmlookup (storage st) (STR ''Victim'') = Some s" - and "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st) (STR ''Victim'')) - SUMM s \ bal \ POS s" - obtains s' where - "loc = sender env + (STR ''.'' + STR ''balance'')" - and "type = STValue (TUInt 256)" - and "fmlookup (storage st') (STR ''Victim'') = Some s'" - and "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st') (STR ''Victim'')) - SUMM s' \ bal \ POS s'" -proof - - from assms(1) obtain v t st'' st''' x - where *: "expr SENDER ep env cd st = Normal ((KValue v, t), st'')" - and **: "ssel (STValue (TUInt 256)) (hash (STR ''balance'') v) [] ep env cd st'' = Normal (x,st''')" - and "st' = st'''" - by (auto split:if_split_asm) - moreover obtain s'' where "v =sender env" - and "t = Value TAddr" - and ***:"fmlookup (storage st'') (STR ''Victim'') = Some s''" - and ****: "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st'') (STR ''Victim'')) - SUMM s'' \ bal \ POS s''" - using securesender[OF * assms(2,3)] by auto - moreover obtain s''' where "x = (hash (STR ''balance'') v, STValue (TUInt 256))" - and "fmlookup (storage st''') (STR ''Victim'') = Some s'''" - and "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st''') (STR ''Victim'')) - SUMM s''' \ bal \ POS s'''" - using securessel[OF ** *** ****] by auto - ultimately show ?thesis using assms(1) that by simp -qed - -lemma securerexp: - assumes "rexp myrexp e\<^sub>p env cd st = Normal ((v, t), st')" - and "fmlookup (denvalue env) (STR ''balance'') = Some (Storage (STMap TAddr (STValue (TUInt 256))), Storeloc STR ''balance'')" - and "fmlookup (storage st) (STR ''Victim'') = Some s" - and "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st) (STR ''Victim'')) - SUMM s \ bal \ POS s" - and "address env = STR ''Victim''" - obtains s' where - "fmlookup (storage st') (address env) = Some s'" - and "v = KValue (accessStorage (TUInt 256) (sender env + (STR ''.'' + STR ''balance'')) s')" - and "t = Value (TUInt 256)" - and "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st') (STR ''Victim'')) - SUMM s' \ bal \ POS s'" -proof - - from assms(1,2) obtain l' t'' st'' s - where *: "ssel (STMap TAddr (STValue (TUInt 256))) (STR ''balance'') [SENDER] e\<^sub>p env cd st = Normal ((l', STValue t''), st'')" - and "fmlookup (storage st'') (address env) = Some s" - and "v = KValue (accessStorage t'' l' s)" - and "t = Value t''" and "st'=st''" - by (simp split:if_split_asm option.split_asm) - moreover obtain s'' where - "fmlookup (storage st'') (STR ''Victim'') = Some s''" - and "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st'') (STR ''Victim'')) - SUMM s'' \ bal \ POS s''" - and "l'=sender env + (STR ''.'' + STR ''balance'')" and "t'' = TUInt 256" using securessel2[OF * assms(3,4)] by blast - ultimately show ?thesis using assms(1,5) that by auto -qed - -lemma securelval: - assumes "expr mylval ep env cd st = Normal((v,t), st')" - and "fmlookup (denvalue env) (STR ''balance'') = Some (Storage (STMap TAddr (STValue (TUInt 256))), Storeloc STR ''balance'')" - and "fmlookup (storage st) (STR ''Victim'') = Some s" - and "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st) (STR ''Victim'')) - SUMM s \ bal \ bal \ 0 \ POS s" - and "address env = STR ''Victim''" - obtains s' where "fmlookup (storage st') (STR ''Victim'') = Some s'" - and "v = KValue (accessStorage (TUInt 256) (sender env + (STR ''.'' + STR ''balance'')) s')" - and "t = Value (TUInt 256)" - and "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st') (STR ''Victim'')) - SUMM s' \ bal \ bal \ 0 \ POS s'" -proof - - define st'' where "st'' = st\gas := gas st - costs\<^sub>e mylval ep env cd st\" - with assms(3,4) have *: "fmlookup (storage st'') (STR ''Victim'') = Some s" - and **: "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st'') (STR ''Victim'')) - SUMM s \ bal \ POS s" by simp+ - - from assms(1) st''_def obtain v' t' st''' where ***: "rexp myrexp ep env cd st'' = Normal ((v, t), st''')" - and "v' = v" - and "t' = t" - and "st''' = st'" - by (simp split:if_split_asm) - - with securerexp[OF *** assms(2) * **] show ?thesis using assms(1,4,5) that by auto -qed - -lemma plus_frame: - assumes "expr (PLUS (LVAL (Ref (STR ''balance'') [SENDER])) VALUE) ep env cd st = Normal (kv, st')" - and "ReadL\<^sub>i\<^sub>n\<^sub>t (accessStorage (TUInt 256) (sender env + (STR ''.'' + STR ''balance'')) s) + ReadL\<^sub>i\<^sub>n\<^sub>t (svalue env) < 2^256" - and "ReadL\<^sub>i\<^sub>n\<^sub>t (accessStorage (TUInt 256) (sender env + (STR ''.'' + STR ''balance'')) s) + ReadL\<^sub>i\<^sub>n\<^sub>t (svalue env) \ 0" - and "fmlookup (storage st) (STR ''Victim'') = Some s" - and "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st) (STR ''Victim'')) - SUMM s \ bal" - and "fmlookup (denvalue env) (STR ''balance'') = Some (Storage (STMap TAddr (STValue (TUInt 256))), Storeloc STR ''balance'')" - and "address env = (STR ''Victim'')" - shows "kv = (KValue (ShowL\<^sub>i\<^sub>n\<^sub>t (ReadL\<^sub>i\<^sub>n\<^sub>t (accessStorage (TUInt 256) (sender env + (STR ''.'' + STR ''balance'')) s) + ReadL\<^sub>i\<^sub>n\<^sub>t (svalue env))), Value (TUInt 256))" - and "fmlookup (storage st') (STR ''Victim'') = Some s" - and "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st') (STR ''Victim'')) = ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st) (STR ''Victim''))" -proof - - define st0 where "st0 = st\gas := gas st - costs\<^sub>e (PLUS (LVAL (Ref (STR ''balance'') [SENDER])) VALUE) ep env cd st\" - define st1 where "st1 = st0\gas := gas st0 - costs\<^sub>e (LVAL (Ref (STR ''balance'') [SENDER])) ep env cd st0\" - define st2 where "st2 = st1\gas := gas st1 - costs\<^sub>e SENDER ep env cd st1\" - - define st3 where "st3 = st2\gas := gas st2 - costs\<^sub>e VALUE ep env cd st2\" - from assms(1) obtain v1 t1 v2 t2 st'' st''' st'''' v t where - *: "expr (LVAL (Ref (STR ''balance'') [SENDER])) ep env cd st0 = Normal ((KValue v1, Value t1), st'')" - and **: "expr VALUE ep env cd st'' = Normal ((KValue v2, Value t2), st''')" - and ***: "add t1 t2 v1 v2 = Some (v,t)" - and ****: "expr (PLUS (LVAL (Ref (STR ''balance'') [SENDER])) VALUE) ep env cd st = Normal ((KValue v, Value t), st'''')" - using st0_def by (auto simp del: expr.simps simp add:expr.simps(11) split:if_split_asm result.split_asm Stackvalue.split_asm Type.split_asm option.split_asm) - - moreover have "expr (LVAL (Ref (STR ''balance'') [SENDER])) ep env cd st0 = Normal ((KValue (accessStorage (TUInt 256) (sender env + (STR ''.'' + STR ''balance'')) s), Value (TUInt 256)), st'')" - and "st'' = st2" - proof - - from * obtain l' t' s'' where *****: "ssel (STMap TAddr (STValue (TUInt 256))) (STR ''balance'') [SENDER] ep env cd st1 = Normal ((l', STValue t'), st'')" - and ******: "fmlookup (storage st'') (address env) = Some s''" and "v1 = (accessStorage t' l' s'')" and "t' = t1" - using st0_def st1_def assms(4,6) by (auto simp del: accessStorage.simps ssel.simps split:if_split_asm option.split_asm STypes.split_asm result.split_asm) - moreover from ***** have "expr SENDER ep env cd st1 = Normal ((KValue (sender env), Value TAddr), st2)" using st2_def by (simp split:if_split_asm) - with ***** have "st'' = st2" and "l' = sender env + (STR ''.'' + STR ''balance'')" and "t' = TUInt 256" by auto - moreover from ****** `st'' = st2` have "s''=s" using st0_def st1_def st2_def assms(4,7) by auto - ultimately show "expr (LVAL (Ref (STR ''balance'') [SENDER])) ep env cd st0 = Normal ((KValue (accessStorage (TUInt 256) (sender env + (STR ''.'' + STR ''balance'')) s), Value (TUInt 256)), st'')" - and "st'' = st2" using * by (auto split:if_split_asm) - qed - - moreover from ** `st'' = st2` have "expr VALUE ep env cd st2 = Normal ((KValue (svalue env), Value (TUInt 256)), st3)" and "st''' = st3" using st1_def st3_def by (auto split:if_split_asm) - moreover have "add (TUInt 256) (TUInt 256) (accessStorage (TUInt 256) (sender env + (STR ''.'' + STR ''balance'')) s) (svalue env) = Some (ShowL\<^sub>i\<^sub>n\<^sub>t (ReadL\<^sub>i\<^sub>n\<^sub>t (accessStorage (TUInt 256) (sender env + (STR ''.'' + STR ''balance'')) s) + ReadL\<^sub>i\<^sub>n\<^sub>t (svalue env)), TUInt 256)" (is "?LHS = ?RHS") - proof - - have "?LHS = Some (createUInt 256 (\(accessStorage (TUInt 256) (sender env + (STR ''.'' + STR ''balance'')) s)\+ \(svalue env)\), TUInt 256)" using add_def olift.simps(2)[of "(+)" 256 256] by simp - with assms(2,3) show "?LHS = ?RHS" by simp - qed - ultimately have "v= (ShowL\<^sub>i\<^sub>n\<^sub>t (ReadL\<^sub>i\<^sub>n\<^sub>t (accessStorage (TUInt 256) (sender env + (STR ''.'' + STR ''balance'')) s) + ReadL\<^sub>i\<^sub>n\<^sub>t (svalue env)))" and "t = TUInt 256" and "st' = st3" using st0_def assms(1) by (auto split:if_split_asm) - with assms(1) **** have "kv = (KValue (ShowL\<^sub>i\<^sub>n\<^sub>t (ReadL\<^sub>i\<^sub>n\<^sub>t (accessStorage (TUInt 256) (sender env + (STR ''.'' + STR ''balance'')) s) + ReadL\<^sub>i\<^sub>n\<^sub>t (svalue env))), Value (TUInt 256))" using st0_def by simp - moreover from assms(4) st0_def st1_def st2_def st3_def `st' = st3` have "fmlookup (storage st') (STR ''Victim'') = Some s" by simp - moreover from assms(5) st0_def st1_def st2_def st3_def `st' = st3` have "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st') (STR ''Victim'')) - SUMM s \ bal" by simp - moreover have "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st') (STR ''Victim'')) = ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st) (STR ''Victim''))" using st0_def st1_def st2_def st3_def `st' = st3` by simp - ultimately show "kv = (KValue (ShowL\<^sub>i\<^sub>n\<^sub>t (ReadL\<^sub>i\<^sub>n\<^sub>t (accessStorage (TUInt 256) (sender env + (STR ''.'' + STR ''balance'')) s) + ReadL\<^sub>i\<^sub>n\<^sub>t (svalue env))), Value (TUInt 256))" - and "fmlookup (storage st') (STR ''Victim'') = Some s" - and "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st') (STR ''Victim'')) = ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st) (STR ''Victim''))" by auto -qed - - - -lemma deposit_frame: - assumes "stmt deposit ep env cd st = Normal ((), st')" - and "fmlookup (storage st) (STR ''Victim'') = Some s" - and "address env = (STR ''Victim'')" - and "fmlookup (denvalue env) (STR ''balance'') = Some (Storage (STMap TAddr (STValue (TUInt 256))), Storeloc STR ''balance'')" - and "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st) (STR ''Victim'')) - SUMM s \ bal + ReadL\<^sub>i\<^sub>n\<^sub>t (svalue env)" - and "ReadL\<^sub>i\<^sub>n\<^sub>t (accessStorage (TUInt 256) (sender env + (STR ''.'' + STR ''balance'')) s) + ReadL\<^sub>i\<^sub>n\<^sub>t (svalue env) < 2^256" - and "ReadL\<^sub>i\<^sub>n\<^sub>t (accessStorage (TUInt 256) (sender env + (STR ''.'' + STR ''balance'')) s) + ReadL\<^sub>i\<^sub>n\<^sub>t (svalue env) \ 0" - and "POS s" - obtains s' - where "fmlookup (storage st') (STR ''Victim'') = Some s'" - and "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st') (STR ''Victim'')) - SUMM s' \ bal" - and "POS s'" -proof - - define st0 where "st0 = st\gas := gas st - costs deposit ep env cd st\" - - from assms(1) st0_def obtain kv st'' where *: "expr (PLUS (LVAL (Ref (STR ''balance'') [SENDER])) VALUE) ep env cd st0 = Normal (kv, st'')" by (auto simp del: expr.simps split:if_split_asm result.split_asm) - moreover have "fmlookup (storage st0) (STR ''Victim'') = Some s" using st0_def assms(2) by simp - moreover from assms(5) have "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st0) (STR ''Victim'')) - SUMM s \ bal + ReadL\<^sub>i\<^sub>n\<^sub>t (svalue env)" using st0_def by simp - ultimately have **: "kv = (KValue \(\accessStorage (TUInt 256) (sender env + (STR ''.'' + STR ''balance'')) s\::int) + \svalue env\\, Value (TUInt 256))" - and st''_s:"fmlookup (storage st'') STR ''Victim'' = Some s" - and ac: "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st'') (STR ''Victim'')) = ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st0) (STR ''Victim''))" - using plus_frame[OF _ assms(6,7) _ _ assms(4,3), of ep cd st0 kv st''] by auto - - define v where "v= (\accessStorage (TUInt 256) (sender env + (STR ''.'' + STR ''balance'')) s\::int) + \svalue env\" - moreover from * ** assms(1) st0_def obtain rl st''' where ***: "lexp (Ref (STR ''balance'') [SENDER]) ep env cd st'' = Normal (rl, st''')" by (auto simp del:expr.simps lexp.simps accessStorage.simps split:if_split_asm result.split_asm) - moreover from *** assms have "rl = (LStoreloc ((sender env) + (STR ''.'' + STR ''balance'')), Storage (STValue (TUInt 256)))" and st'''_def: "st''' = st''\gas := gas st'' - costs\<^sub>e SENDER ep env cd st''\" - proof - - from *** assms(4) obtain l' t' where - "fmlookup (denvalue env) (STR ''balance'') = Some (Storage (STMap TAddr (STValue (TUInt 256))), Storeloc (STR ''balance''))" - and *:"ssel (STMap TAddr (STValue (TUInt 256))) (STR ''balance'') [SENDER] ep env cd st'' = Normal ((l',t'), st''')" - and "rl = (LStoreloc l', Storage t')" - by (auto simp del: ssel.simps split:if_split_asm option.split_asm result.split_asm) - moreover from * have "ssel (STMap TAddr (STValue (TUInt 256))) (STR ''balance'') [SENDER] ep env cd st'' = Normal ((((sender env) + (STR ''.'' + STR ''balance'')), STValue (TUInt 256)), st''\gas := gas st'' - costs\<^sub>e SENDER ep env cd st''\)" by (simp split:if_split_asm) - ultimately show "rl = (LStoreloc ((sender env) + (STR ''.'' + STR ''balance'')), Storage (STValue (TUInt 256)))" and st'''_def: "st''' = st''\gas := gas st'' - costs\<^sub>e SENDER ep env cd st''\" by auto - qed - moreover have "Valuetypes.convert (TUInt 256) (TUInt 256) (ShowL\<^sub>i\<^sub>n\<^sub>t v) = Some (ShowL\<^sub>i\<^sub>n\<^sub>t v, TUInt 256)" by simp - - moreover from st''_s st'''_def have s'''_s: "fmlookup (storage st''') (STR ''Victim'') = Some s" by simp - ultimately have ****:"st' = st'''\storage := fmupd (STR ''Victim'') (fmupd ((sender env) + (STR ''.'' + STR ''balance'')) (ShowL\<^sub>i\<^sub>n\<^sub>t v) s) (storage st''')\" - using assms(1) * ** st0_def assms(3) by (auto simp del:expr.simps lexp.simps accessStorage.simps split:if_split_asm) - - moreover define s' where "s' = (fmupd ((sender env) + (STR ''.'' + STR ''balance'')) (ShowL\<^sub>i\<^sub>n\<^sub>t v) s)" - ultimately have "fmlookup (storage st') (STR ''Victim'') = Some s'" - and *****:"fmlookup s' ((sender env) + (STR ''.'' + STR ''balance'')) = Some (ShowL\<^sub>i\<^sub>n\<^sub>t v)" by simp_all - - moreover have "SUMM s' = SUMM s + \svalue env\" - proof - - have s1: "SUMM s = (\(ad,x)|fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env. ReadL\<^sub>i\<^sub>n\<^sub>t x) + ReadL\<^sub>i\<^sub>n\<^sub>t (accessStorage (TUInt 256) (sender env + (STR ''.'' + STR ''balance'')) s)" - proof (cases "fmlookup s (sender env + (STR ''.'' + STR ''balance'')) = None") - case True - then have "accessStorage (TUInt 256) (sender env + (STR ''.'' + STR ''balance'')) s = ShowL\<^sub>i\<^sub>n\<^sub>t 0" by simp - moreover have "{(ad,x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x} = {(ad,x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env}" - proof - show "{(ad, x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x} \ {(ad, x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env}" - proof - fix x - assume "x \ {(ad, x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x}" - then show "x \ {(ad, x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env}" using True by auto - qed - next - show "{(ad, x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env} \ {(ad, x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x }" - proof - fix x - assume "x \ {(ad, x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env}" - then show "x \ {(ad, x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x}" using True by auto - qed - qed - then have "SUMM s = (\(ad,x)|fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env. ReadL\<^sub>i\<^sub>n\<^sub>t x)" by simp - ultimately show ?thesis using Read_ShowL_id by simp - next - case False - then obtain val where val_def: "fmlookup s (sender env + (STR ''.'' + STR ''balance'')) = Some val" by auto - - have "inj_on (\(ad, x). (ad + (STR ''.'' + STR ''balance''), x)) {(ad, x). (fmlookup s \ (\ad. ad + (STR ''.'' + STR ''balance''))) ad = Some x}" using balance_inj by simp - then have "finite {(ad, x). (fmlookup s \ (\ad. ad + (STR ''.'' + STR ''balance''))) ad = Some x}" using fmlookup_finite[of "\ad. (ad + (STR ''.'' + STR ''balance''))" s] by simp - then have sum1: "finite ({(ad,x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env})" using finite_subset[of "{(ad, x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env}" "{(ad, x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x}"] by auto - moreover have sum2: "(sender env,val) \ {(ad,x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env}" by simp - moreover from sum1 x1 val_def have "insert (sender env,val) {(ad, x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env} = {(ad, x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x}" by auto - ultimately show ?thesis using sum.insert[OF sum1 sum2, of "\(ad,x). ReadL\<^sub>i\<^sub>n\<^sub>t x"] val_def by simp - qed - moreover have s2: "SUMM s' = (\(ad,x)|fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env. ReadL\<^sub>i\<^sub>n\<^sub>t x) + v" - proof - - have "inj_on (\(ad, x). (ad + (STR ''.'' + STR ''balance''), x)) {(ad, x). (fmlookup s' \ (\ad. ad + (STR ''.'' + STR ''balance''))) ad = Some x}" using balance_inj by simp - then have "finite {(ad, x). (fmlookup s' \ (\ad. ad + (STR ''.'' + STR ''balance''))) ad = Some x}" using fmlookup_finite[of "\ad. (ad + (STR ''.'' + STR ''balance''))" s'] by simp - then have sum1: "finite ({(ad,x). fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env})" using finite_subset[of "{(ad, x). fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env}" "{(ad, x). fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x}"] by auto - moreover have sum2: "(sender env,ShowL\<^sub>i\<^sub>n\<^sub>t v) \ {(ad,x). fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env}" by simp - moreover from ***** have "insert (sender env,ShowL\<^sub>i\<^sub>n\<^sub>t v) {(ad, x). fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env} = {(ad, x). fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x}" by auto - ultimately show ?thesis using sum.insert[OF sum1 sum2, of "\(ad,x). ReadL\<^sub>i\<^sub>n\<^sub>t x"] Read_ShowL_id by simp - qed - moreover have s3: "(\(ad,x)|fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env. ReadL\<^sub>i\<^sub>n\<^sub>t x) - =(\(ad,x)|fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env. ReadL\<^sub>i\<^sub>n\<^sub>t x)" - proof - - have "{(ad,x). fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env} = {(ad,x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env}" - proof - show "{(ad, x). fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env} - \ {(ad, x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env}" - proof - fix xx - assume "xx \ {(ad, x). fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env}" - then obtain ad x where "xx = (ad,x)" and "fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x" and "ad \ sender env" by auto - then have "fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x" using s'_def String_Cancel[of ad "(STR ''.'' + STR ''balance'')" "sender env"] by (simp split:if_split_asm) - with `ad \ sender env` `xx = (ad,x)` show "xx \ {(ad, x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env}" by simp - qed - next - show "{(ad, x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env} - \ {(ad, x). fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env}" - proof - fix xx - assume "xx \ {(ad, x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env}" - then obtain ad x where "xx = (ad,x)" and "fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x" and "ad \ sender env" by auto - then have "fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x" using s'_def String_Cancel[of ad "(STR ''.'' + STR ''balance'')" "sender env"] by (auto split:if_split_asm) - with `ad \ sender env` `xx = (ad,x)` show "xx \ {(ad, x). fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env}" by simp - qed - qed - thus ?thesis by simp - qed - moreover from s'_def v_def have "ReadL\<^sub>i\<^sub>n\<^sub>t (accessStorage (TUInt 256) (sender env + (STR ''.'' + STR ''balance'')) s') = ReadL\<^sub>i\<^sub>n\<^sub>t (accessStorage (TUInt 256) (sender env + (STR ''.'' + STR ''balance'')) s) + \svalue env\" using Read_ShowL_id by (simp split:option.split_asm) - ultimately have "SUMM s' = SUMM s + \svalue env\" - proof - - from s2 have "SUMM s' = (\(ad,x)|fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env. ReadL\<^sub>i\<^sub>n\<^sub>t x) + v" by simp - also from s3 have "\ = (\(ad,x)|fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env. ReadL\<^sub>i\<^sub>n\<^sub>t x) + v" by simp - also from s1 have "\ = SUMM s - ReadL\<^sub>i\<^sub>n\<^sub>t (accessStorage (TUInt 256) (sender env + (STR ''.'' + STR ''balance'')) s) + v" by simp - finally show ?thesis using v_def by simp - qed - then show ?thesis by simp - qed - moreover have "POS s'" - proof (rule allI[OF allI]) - fix ad xa - show "fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some xa \ 0 \ (\xa\::int)" - proof - assume a1: "fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some xa" - show "0 \ (\xa\::int)" - proof (cases "ad = sender env") - case True - then show ?thesis using s'_def assms(7) Read_ShowL_id a1 v_def by auto - next - case False - then show ?thesis using s'_def assms(7,8) using Read_ShowL_id a1 v_def by (auto split:if_split_asm) - qed - qed - qed - moreover have "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st') (STR ''Victim'')) = ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st) (STR ''Victim''))" using **** ac st0_def st'''_def by simp - ultimately show ?thesis using that assms(5) by simp -qed - - -lemma secure: - "address ev1 \ (STR ''Victim'') \ fmlookup ep1 (STR ''Victim'') = Some (victim, SKIP) \ (\rv1 st1' bal. frame bal st1 \ msel c1 t1 l1 xe1 ep1 ev1 cd1 st1 = Normal (rv1, st1') \ frame bal st1')" - "address ev2 \ (STR ''Victim'') \ fmlookup ep2 (STR ''Victim'') = Some (victim, SKIP) \ (\rv2 st2' bal. frame bal st2 \ ssel t2 l2 xe2 ep2 ev2 cd2 st2 = Normal (rv2, st2') \ frame bal st2')" - "address ev5 \ (STR ''Victim'') \ fmlookup ep5 (STR ''Victim'') = Some (victim, SKIP) \ (\rv3 st5' bal. frame bal st5 \ lexp l5 ep5 ev5 cd5 st5 = Normal (rv3, st5') \ frame bal st5')" - "address ev4 \ (STR ''Victim'') \ fmlookup ep4 (STR ''Victim'') = Some (victim, SKIP) \ (\rv4 st4' bal. frame bal st4 \ expr e4 ep4 ev4 cd4 st4 = Normal (rv4, st4') \ frame bal st4')" - "address lev \ (STR ''Victim'') \ fmlookup lep (STR ''Victim'') = Some (victim, SKIP) \ (\ev cd st st' bal. load lcp lis lxs lep lev0 lcd0 lst0 lev lcd lst = Normal ((ev, cd, st), st') \ (frame bal lst0 \ frame bal st) \ (frame bal lst \ frame bal st') \ address lev0 = address ev \ sender lev0 = sender ev \ svalue lev0 = svalue ev)" - "address ev3 \ (STR ''Victim'') \ fmlookup ep3 (STR ''Victim'') = Some (victim, SKIP) \ (\rv3 st3' bal. frame bal st3 \ rexp l3 ep3 ev3 cd3 st3 = Normal (rv3, st3') \ frame bal st3')" - "(fmlookup ep6 (STR ''Victim'') = Some (victim, SKIP) \ - (\st6'. stmt s6 ep6 ev6 cd6 st6 = Normal((), st6') \ - ((address ev6 \ (STR ''Victim'') \ (\bal. frame bal st6 \ frame bal st6')) - \ (address ev6 = (STR ''Victim'') \ - (\s val bal x. s6 = transfer - \ INV st6 s (SUMM s + ReadL\<^sub>i\<^sub>n\<^sub>t val) bal \ POS s - \ fmlookup (denvalue ev6) (STR ''bal'') = Some (Value (TUInt 256), Stackloc x) - \ accessStore x (stack st6) = Some (KValue val) - \ sender ev6 \ address ev6 - \ (\s'. fmlookup (storage st6') (STR ''Victim'') = Some s' - \ ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st6') (STR ''Victim'')) - (SUMM s') \ bal \ bal \ 0 \ POS s')) \ - (\s bal x. s6 = comp - \ INV st6 s (SUMM s) bal \ POS s - \ fmlookup (denvalue ev6) (STR ''bal'') = Some (Value (TUInt 256), Stackloc x) - \ fmlookup (denvalue ev6) (STR ''balance'') = Some (Storage (STMap TAddr (STValue (TUInt 256))), Storeloc STR ''balance'') - \ accessStore x (stack st6) = Some (KValue (accessStorage (TUInt 256) (sender ev6 + (STR ''.'' + STR ''balance'')) s)) - \ sender ev6 \ address ev6 - \ (\s'. fmlookup (storage st6') (STR ''Victim'') = Some s' - \ ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st6') (STR ''Victim'')) - (SUMM s') \ bal \ bal \ 0 \ POS s')) \ - (\s bal. s6 = keep - \ INV st6 s (SUMM s) bal \ POS s - \ fmlookup (denvalue ev6) (STR ''balance'') = Some (Storage (STMap TAddr (STValue (TUInt 256))), Storeloc STR ''balance'') - \ sender ev6 \ address ev6 - \ (\s'. fmlookup (storage st6') (STR ''Victim'') = Some s' - \ ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st6') (STR ''Victim'')) - (SUMM s') \ bal \ bal \ 0 \ POS s')) -))))" -proof (induct rule: msel_ssel_lexp_expr_load_rexp_stmt.induct -[where ?P1.0="\c1 t1 l1 xe1 ep1 ev1 cd1 st1. address ev1 \ (STR ''Victim'') \ fmlookup ep1 (STR ''Victim'') = Some (victim, SKIP) \ (\rv1 st1' bal. frame bal st1 \ msel c1 t1 l1 xe1 ep1 ev1 cd1 st1 = Normal (rv1, st1') \ frame bal st1')" - and ?P2.0="\t2 l2 xe2 ep2 ev2 cd2 st2. address ev2 \ (STR ''Victim'') \ fmlookup ep2 (STR ''Victim'') = Some (victim, SKIP) \ (\rv2 st2' bal. frame bal st2 \ ssel t2 l2 xe2 ep2 ev2 cd2 st2 = Normal (rv2, st2') \ frame bal st2')" - and ?P3.0="\l5 ep5 ev5 cd5 st5. address ev5 \ (STR ''Victim'') \ fmlookup ep5 (STR ''Victim'') = Some (victim, SKIP) \ (\rv5 st5' bal. frame bal st5 \ lexp l5 ep5 ev5 cd5 st5 = Normal (rv5, st5') \ frame bal st5')" - and ?P4.0="\e4 ep4 ev4 cd4 st4. address ev4 \ (STR ''Victim'') \ fmlookup ep4 (STR ''Victim'') = Some (victim, SKIP) \ (\rv4 st4' bal. frame bal st4 \ expr e4 ep4 ev4 cd4 st4 = Normal (rv4, st4') \ frame bal st4')" - and ?P5.0="\lcp lis lxs lep lev0 lcd0 lst0 lev lcd lst. address lev \ (STR ''Victim'') \ fmlookup lep (STR ''Victim'') = Some (victim, SKIP) \ (\ev cd st st' bal. load lcp lis lxs lep lev0 lcd0 lst0 lev lcd lst = Normal ((ev, cd, st), st') \ (frame bal lst0 \ frame bal st) \ (frame bal lst \ frame bal st') \ address lev0 = address ev \ sender lev0 = sender ev \ svalue lev0 = svalue ev)" - and ?P6.0="\l3 ep3 ev3 cd3 st3. address ev3 \ (STR ''Victim'') \ fmlookup ep3 (STR ''Victim'') = Some (victim, SKIP) \ (\rv3 st3' bal. frame bal st3 \ rexp l3 ep3 ev3 cd3 st3 = Normal (rv3, st3') \ frame bal st3')" - and ?P7.0="\s6 ep6 ev6 cd6 st6. -(fmlookup ep6 (STR ''Victim'') = Some (victim, SKIP) \ - (\st6'. stmt s6 ep6 ev6 cd6 st6 = Normal((), st6') \ - ((address ev6 \ (STR ''Victim'') \ (\bal. frame bal st6 \ frame bal st6')) - \ (address ev6 = (STR ''Victim'') \ - (\s val bal x. s6 = transfer - \ INV st6 s (SUMM s + ReadL\<^sub>i\<^sub>n\<^sub>t val) bal \ POS s - \ fmlookup (denvalue ev6) (STR ''bal'') = Some (Value (TUInt 256), Stackloc x) - \ accessStore x (stack st6) = Some (KValue val) - \ sender ev6 \ address ev6 - \ (\s'. fmlookup (storage st6') (STR ''Victim'') = Some s' - \ ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st6') (STR ''Victim'')) - (SUMM s') \ bal \ bal \ 0 \ POS s')) \ - (\s bal x. s6 = comp - \ INV st6 s (SUMM s) bal \ POS s - \ fmlookup (denvalue ev6) (STR ''bal'') = Some (Value (TUInt 256), Stackloc x) - \ fmlookup (denvalue ev6) (STR ''balance'') = Some (Storage (STMap TAddr (STValue (TUInt 256))), Storeloc STR ''balance'') - \ accessStore x (stack st6) = Some (KValue (accessStorage (TUInt 256) (sender ev6 + (STR ''.'' + STR ''balance'')) s)) - \ sender ev6 \ address ev6 - \ (\s'. fmlookup (storage st6') (STR ''Victim'') = Some s' - \ ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st6') (STR ''Victim'')) - (SUMM s') \ bal \ bal \ 0 \ POS s')) \ - (\s bal. s6 = keep - \ INV st6 s (SUMM s) bal \ POS s - \ fmlookup (denvalue ev6) (STR ''balance'') = Some (Storage (STMap TAddr (STValue (TUInt 256))), Storeloc STR ''balance'') - \ sender ev6 \ address ev6 - \ (\s'. fmlookup (storage st6') (STR ''Victim'') = Some s' - \ ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st6') (STR ''Victim'')) - (SUMM s') \ bal \ bal \ 0 \ POS s')) -))))" -]) - case (1 uu uv uw ux uy uz st) - then show ?case by simp -next - case (2 va vb vc vd ve vf vg st) - then show ?case by simp -next - case (3 vh al t loc x e\<^sub>p env cd st) - show ?case (is "_ \ ?RHS") - proof - assume asm: "address env \ (STR ''Victim'') \ fmlookup e\<^sub>p (STR ''Victim'') = Some (victim, SKIP)" - show ?RHS - proof (rule allI[OF allI[OF allI[OF impI]]]) - fix rv1 and st' and bal - assume *: "frame bal st \ msel vh (MTArray al t) loc [x] e\<^sub>p env cd st = Normal (rv1, st')" - moreover from * obtain v4 t4 st4' where **: "expr x e\<^sub>p env cd st = Normal ((v4, t4), st4')" by (auto split: result.split_asm) - moreover from * ** have "frame bal st4'" using 3(1) asm by (auto split:if_split_asm) - ultimately show "frame bal st'" by (simp split:Stackvalue.split_asm Type.split_asm if_split_asm) - qed - qed -next - case (4 mm al t loc x y ys e\<^sub>p env cd st) - show ?case (is "_ \ ?RHS") - proof - assume asm: "address env \ (STR ''Victim'') \ fmlookup e\<^sub>p (STR ''Victim'') = Some (victim, SKIP)" - show ?RHS - proof (rule allI[OF allI[OF allI[OF impI]]]) - fix rv1 and st' and bal - assume *: "frame bal st \ msel mm (MTArray al t) loc (x # y # ys) e\<^sub>p env cd st = Normal (rv1, st')" - moreover from * obtain v4 t4 st'' where **: "expr x e\<^sub>p env cd st = Normal ((KValue v4, Value t4), st'')" by (auto split: result.split_asm Stackvalue.split_asm Type.split_asm) - moreover from * ** have f1: "frame bal st''" using 4(1) asm by (auto split:if_split_asm) - moreover from * ** have ***: "Valuetypes.less t4 (TUInt 256) v4 \al\ = Some (\True\, TBool)" by (auto split: result.split_asm Stackvalue.split_asm Type.split_asm if_split_asm) - moreover from * ** *** obtain vb st''' where ****: "(applyf (\st. if mm then memory st else cd) st'') = Normal (vb, st''')" - and f2: "frame bal st'''" using f1 by (simp split:Stackvalue.split_asm Type.split_asm if_split_asm) - moreover from * ** *** **** obtain ll where *****: "accessStore (hash loc v4) vb = Some (MPointer ll)" - by (simp split: Type.split_asm if_split_asm option.split_asm Memoryvalue.split_asm) - moreover from * ** *** **** ***** obtain l1' st'''' where ******: "msel mm t ll (y # ys) e\<^sub>p env cd st''' = Normal (l1', st'''')" - by (simp split: Type.split_asm if_split_asm option.split_asm Memoryvalue.split_asm) - ultimately have "st' = st''''" by simp - moreover have x1: "\rv1' st1' bal. (frame bal st''') \ local.msel mm t ll (y # ys) e\<^sub>p env cd st''' = Normal (rv1', st1') \ frame bal st1'" using 4(2)[OF ** _ _ _ *** _ *****] **** asm apply safe by auto - ultimately show "frame bal st'" using f2 ****** by blast - qed - qed -next - case (5 tp loc vi vj vk st) - then show ?case by simp -next - case (6 vl vm vn vo vp vq vr st) - then show ?case by simp -next - case (7 al t loc x xs e\<^sub>p env cd st) - show ?case (is "_ \ ?RHS") - proof - assume asm: "address env \ (STR ''Victim'') \ fmlookup e\<^sub>p (STR ''Victim'') = Some (victim, SKIP)" - show ?RHS - proof (rule allI[OF allI[OF allI[OF impI]]]) - fix rv1 and st' and bal - assume *: "frame bal st \ ssel (STArray al t) loc (x # xs) e\<^sub>p env cd st = Normal (rv1, st')" - moreover from * obtain v4 t4 st4' where **: "expr x e\<^sub>p env cd st = Normal ((KValue v4, Value t4), st4')" by (auto split: result.split_asm Stackvalue.split_asm Type.split_asm) - moreover from * ** have f1: "frame bal st4'" using 7(1) asm by (auto split:if_split_asm) - moreover from * ** have ***: "Valuetypes.less t4 (TUInt 256) v4 \al\ = Some (\True\, TBool)" by (auto split: result.split_asm Stackvalue.split_asm Type.split_asm if_split_asm) - moreover from * ** *** obtain l1' st''' where ****: "ssel t (hash loc v4) xs e\<^sub>p env cd st4' = Normal (l1', st''')" by (simp split: Type.split_asm if_split_asm option.split_asm Memoryvalue.split_asm) - ultimately have "st' = st'''" by simp - moreover have "\rv' st2' bal. (frame bal st4') \ ssel t (hash loc v4) xs e\<^sub>p env cd st4' = Normal (rv', st2') \ frame bal st2'" using 7(2)[OF ** _ _ _ ***] asm apply safe by auto - ultimately show "frame bal st'" using f1 **** by blast - qed - qed -next - case (8 vs t loc x xs e\<^sub>p env cd st) - show ?case (is "_ \ ?RHS") - proof - assume asm: "address env \ (STR ''Victim'') \ fmlookup e\<^sub>p (STR ''Victim'') = Some (victim, SKIP)" - show ?RHS - proof (rule allI[OF allI[OF allI[OF impI]]]) - fix rv1 and st' and bal - assume *: "frame bal st \ ssel (STMap vs t) loc (x # xs) e\<^sub>p env cd st = Normal (rv1, st')" - moreover from * obtain v4 t4 st4' where **: "expr x e\<^sub>p env cd st = Normal ((KValue v4, t4), st4')" by (auto split: result.split_asm Stackvalue.split_asm) - moreover from * ** have ***: "frame bal st4'" using 8(1) asm by (auto split:if_split_asm) - moreover from * ** *** obtain l1' st''' where ****:"ssel t (hash loc v4) xs e\<^sub>p env cd st4' = Normal (l1', st''')" by simp - moreover from *** **** have "frame bal st'''" using 8(2)[OF **,of "KValue v4" t4 v4] asm by blast - ultimately show "frame bal st'" by (simp split:Stackvalue.split_asm) - qed - qed -next - case (9 i vt e vu st) - then show ?case by (auto split:option.split_asm result.split_asm Denvalue.split_asm) -next - case (10 i r e\<^sub>p e cd st) - show ?case (is "_ \ ?RHS") - proof - assume asm: "address e \ (STR ''Victim'') \ fmlookup e\<^sub>p (STR ''Victim'') = Some (victim, SKIP)" - show ?RHS - proof (rule allI[OF allI[OF allI[OF impI]]]) - fix rv1 and st' and bal - assume *: "frame bal st \ lexp (Ref i r) e\<^sub>p e cd st = Normal (rv1, st')" - show "frame bal st'" - proof (cases "fmlookup (denvalue e) i") - case None - with * show ?thesis by simp - next - case (Some a) - then show ?thesis - proof (cases a) - case (Pair tp b) - then show ?thesis - proof (cases b) - case (Stackloc l') - then show ?thesis - proof (cases "accessStore l' (stack st)") - case None - with * show ?thesis using Some Pair Stackloc by simp - next - case s2: (Some k) - then show ?thesis - proof (cases k) - case (KValue x1) - with * show ?thesis using Some Pair Stackloc s2 by simp - next - case (KCDptr x2) - with * show ?thesis using Some Pair Stackloc s2 by simp - next - case (KMemptr l'') - then show ?thesis - proof (cases tp) - case (Value x1) - with * show ?thesis using Some Pair Stackloc s2 KMemptr by simp - next - case (Calldata x2) - with * show ?thesis using Some Pair Stackloc s2 KMemptr by simp - next - case (Memory x3) - with * Some Pair Stackloc KMemptr s2 obtain l1' t1' where "msel True x3 l'' r e\<^sub>p e cd st = Normal ((l1', t1'), st')" by (auto split: result.split_asm) - with * 10(1)[OF Some Pair Stackloc _ _ KMemptr, of "Some k" st x3] show ?thesis using s2 Memory asm by auto - next - case (Storage x4) - with * show ?thesis using Some Pair Stackloc s2 KMemptr by simp - qed - next - case (KStoptr l'') - then show ?thesis - proof (cases tp) - case (Value x1) - with * show ?thesis using Some Pair Stackloc s2 KStoptr by simp - next - case (Calldata x2) - with * show ?thesis using Some Pair Stackloc s2 KStoptr by simp - next - case (Memory x3) - with * show ?thesis using Some Pair Stackloc s2 KStoptr by simp - next - case (Storage x4) - with * Some Pair Stackloc KStoptr s2 obtain l1' t1' where "ssel x4 l'' r e\<^sub>p e cd st = Normal ((l1', t1'), st')" by (auto split: result.split_asm) - with * 10(2)[OF Some Pair Stackloc _ _ KStoptr, of "Some k" st x4] show ?thesis using s2 Storage asm by auto - qed - qed - qed - next - case (Storeloc l'') - then show ?thesis - proof (cases tp) - case (Value x1) - with * show ?thesis using Some Pair Storeloc by simp - next - case (Calldata x2) - with * show ?thesis using Some Pair Storeloc by simp - next - case (Memory x3) - with * show ?thesis using Some Pair Storeloc by simp - next - case (Storage x4) - with * Some Pair Storeloc obtain l1' t1' where "ssel x4 l'' r e\<^sub>p e cd st = Normal ((l1', t1'), st')" by (auto split: result.split_asm) - with * 10(3)[OF Some Pair Storeloc Storage] asm show ?thesis by auto - qed - qed - qed - qed - qed - qed -next - case (11 b x e\<^sub>p e cd st) - then show ?case by (simp add:frame_def) -next - case (12 b x e\<^sub>p e cd st) - then show ?case by (simp add:frame_def) -next - case (13 ad e\<^sub>p e cd st) - then show ?case by (simp add:frame_def) -next - case (14 ad e\<^sub>p e cd st) - show ?case (is "_ \ ?RHS") - proof - assume asm: "address e \ (STR ''Victim'') \ fmlookup e\<^sub>p (STR ''Victim'') = Some (victim, SKIP)" - show ?RHS - proof (rule allI[OF allI[OF allI[OF impI]]]) - fix rv1 and st' and bal - assume *: "frame bal st \ expr (BALANCE ad) e\<^sub>p e cd st = Normal (rv1, st')" - moreover from * obtain adv st'' where **:"expr ad e\<^sub>p e cd (st\gas:=gas st - (costs\<^sub>e (BALANCE ad) e\<^sub>p e cd st)\) = Normal ((KValue adv, Value TAddr),st'')" by (auto split:if_split_asm result.split_asm Stackvalue.split_asm Types.split_asm Type.split_asm) - with * ** have "frame bal st''" using 14(1) asm by (auto simp add:frame_def split:if_split_asm) - moreover from * ** have "st' = st''" by (simp split:if_split_asm) - ultimately show "frame bal st'" by simp - qed - qed -next - case (15 e\<^sub>p e cd st) - then show ?case by (simp add:frame_def) -next - case (16 e\<^sub>p e cd st) - then show ?case by (simp add:frame_def) -next - case (17 e\<^sub>p e cd st) - then show ?case by (simp add:frame_def) -next - case (18 e\<^sub>p e cd st) - then show ?case by (simp add:frame_def) -next - case (19 e\<^sub>p e cd st) - then show ?case by (simp add:frame_def) -next - case (20 x e\<^sub>p e cd st) - show ?case (is "_ \ ?RHS") - proof - assume asm: "address e \ (STR ''Victim'') \ fmlookup e\<^sub>p (STR ''Victim'') = Some (victim, SKIP)" - show ?RHS - proof (rule allI[OF allI[OF allI[OF impI]]]) - fix rv1 and st' and bal - assume *: "frame bal st \ expr (NOT x) e\<^sub>p e cd st = Normal (rv1, st')" - then have f1: "frame bal (st\gas:=gas st - (costs\<^sub>e (NOT x) e\<^sub>p e cd st)\)" by (simp add:frame_def) - moreover from * obtain v t st'' where **: "expr x e\<^sub>p e cd (st\gas:=gas st - (costs\<^sub>e (NOT x) e\<^sub>p e cd st)\) = Normal ((KValue v, Value t), st'')" - by (auto split:if_split_asm result.split_asm prod.split_asm Stackvalue.split_asm Type.split_asm) - moreover from * ** have ***: "frame bal st''" using 20(1) asm by (auto simp add:frame_def split:if_split_asm) - show "frame bal st'" - proof (cases "v = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True") - case True - with * ** *** obtain x st''' where "expr FALSE e\<^sub>p e cd st'' = Normal (x, st''')" - and "frame bal st'''" by (auto simp add:frame_def split:if_split_asm) - with * ** *** True show ?thesis by (auto split: if_split_asm) - next - case False - with * ** *** obtain x st''' where "expr TRUE e\<^sub>p e cd st'' = Normal (x, st''')" - and "frame bal st'''" by (auto simp add:frame_def split:if_split_asm) - with * ** *** False show ?thesis by (auto split: if_split_asm) - qed - qed - qed -next - case (21 e1 e2 e\<^sub>p e cd st) - show ?case (is "_ \ ?RHS") - proof - assume asm: "address e \ (STR ''Victim'') \ fmlookup e\<^sub>p (STR ''Victim'') = Some (victim, SKIP)" - show ?RHS - proof (rule allI[OF allI[OF allI[OF impI]]]) - fix rv1 and st' and bal - assume *: "frame bal st \ expr (PLUS e1 e2) e\<^sub>p e cd st = Normal (rv1, st')" - moreover from * obtain v1 t1 st'' where **: "expr e1 e\<^sub>p e cd (st\gas:=gas st - (costs\<^sub>e (PLUS e1 e2) e\<^sub>p e cd st)\) = Normal ((KValue v1, Value t1), st'')" - by (auto split:if_split_asm result.split_asm prod.split_asm Stackvalue.split_asm Type.split_asm) - moreover from * ** have ***: "frame bal st''" using 21(1) asm by (auto simp add:frame_def split:if_split_asm) - moreover from * ** *** obtain v2 t2 st''' where ****: "expr e2 e\<^sub>p e cd st'' = Normal ((KValue v2, Value t2), st''')" - by (auto split:if_split_asm result.split_asm prod.split_asm Stackvalue.split_asm Type.split_asm) - moreover from * ** *** **** have "frame bal st'''" using 21(2)[OF _ **] asm by (auto split:if_split_asm) - moreover from * ** **** obtain v t where "add t1 t2 v1 v2 = Some (v, t)" by (auto split:if_split_asm option.split_asm) - ultimately show "frame bal st'" by (auto split:if_split_asm) - qed - qed -next - case (22 e1 e2 e\<^sub>p e cd st) - show ?case (is "_ \ ?RHS") - proof - assume asm: "address e \ (STR ''Victim'') \ fmlookup e\<^sub>p (STR ''Victim'') = Some (victim, SKIP)" - show ?RHS - proof (rule allI[OF allI[OF allI[OF impI]]]) - fix rv1 and st' and bal - assume *: "frame bal st \ expr (MINUS e1 e2) e\<^sub>p e cd st = Normal (rv1, st')" - moreover from * obtain v1 t1 st'' where **: "expr e1 e\<^sub>p e cd (st\gas:=gas st - (costs\<^sub>e (MINUS e1 e2) e\<^sub>p e cd st)\) = Normal ((KValue v1, Value t1), st'')" - by (auto split:if_split_asm result.split_asm prod.split_asm Stackvalue.split_asm Type.split_asm) - moreover from * ** have ***: "frame bal st''" using 22(1) asm by (auto simp add:frame_def split:if_split_asm) - moreover from * ** *** obtain v2 t2 st''' where ****: "expr e2 e\<^sub>p e cd st'' = Normal ((KValue v2, Value t2), st''')" - by (auto split:if_split_asm result.split_asm prod.split_asm Stackvalue.split_asm Type.split_asm) - moreover from * ** *** **** have "frame bal st'''" using 22(2)[OF _ **] asm by (auto split:if_split_asm) - moreover from * ** **** obtain v t where "sub t1 t2 v1 v2 = Some (v, t)" by (auto split:if_split_asm option.split_asm) - ultimately show "frame bal st'" by (auto split:if_split_asm) - qed - qed -next - case (23 e1 e2 e\<^sub>p e cd st) - show ?case (is "_ \ ?RHS") - proof - assume asm: "address e \ (STR ''Victim'') \ fmlookup e\<^sub>p (STR ''Victim'') = Some (victim, SKIP)" - show ?RHS - proof (rule allI[OF allI[OF allI[OF impI]]]) - fix rv1 and st' and bal - assume *: "frame bal st \ expr (LESS e1 e2) e\<^sub>p e cd st = Normal (rv1, st')" - moreover from * obtain v1 t1 st'' where **: "expr e1 e\<^sub>p e cd (st\gas:=gas st - (costs\<^sub>e (LESS e1 e2) e\<^sub>p e cd st)\) = Normal ((KValue v1, Value t1), st'')" - by (auto split:if_split_asm result.split_asm prod.split_asm Stackvalue.split_asm Type.split_asm) - moreover from * ** have ***: "frame bal st''" using 23(1) asm by (auto simp add:frame_def split:if_split_asm) - moreover from * ** *** obtain v2 t2 st''' where ****: "expr e2 e\<^sub>p e cd st'' = Normal ((KValue v2, Value t2), st''')" - by (auto split:if_split_asm result.split_asm prod.split_asm Stackvalue.split_asm Type.split_asm) - moreover from * ** *** **** have "frame bal st'''" using 23(2)[OF _ **] asm by (auto split:if_split_asm) - moreover from * ** **** obtain v t where "Valuetypes.less t1 t2 v1 v2 = Some (v, t)" by (auto split:if_split_asm option.split_asm) - ultimately show "frame bal st'" by (auto split:if_split_asm) - qed - qed -next - case (24 e1 e2 e\<^sub>p e cd st) - show ?case (is "_ \ ?RHS") - proof - assume asm: "address e \ (STR ''Victim'') \ fmlookup e\<^sub>p (STR ''Victim'') = Some (victim, SKIP)" - show ?RHS - proof (rule allI[OF allI[OF allI[OF impI]]]) - fix rv1 and st' and bal - assume *: "frame bal st \ expr (EQUAL e1 e2) e\<^sub>p e cd st = Normal (rv1, st')" - moreover from * obtain v1 t1 st'' where **: "expr e1 e\<^sub>p e cd (st\gas:=gas st - (costs\<^sub>e (EQUAL e1 e2) e\<^sub>p e cd st)\) = Normal ((KValue v1, Value t1), st'')" - by (auto split:if_split_asm result.split_asm prod.split_asm Stackvalue.split_asm Type.split_asm) - moreover from * ** have ***: "frame bal st''" using 24(1) asm by (auto simp add:frame_def split:if_split_asm) - moreover from * ** *** obtain v2 t2 st''' where ****: "expr e2 e\<^sub>p e cd st'' = Normal ((KValue v2, Value t2), st''')" - by (auto split:if_split_asm result.split_asm prod.split_asm Stackvalue.split_asm Type.split_asm) - moreover from * ** *** **** have "frame bal st'''" using 24(2)[OF _ **] asm by (auto split:if_split_asm) - moreover from * ** **** obtain v t where "Valuetypes.equal t1 t2 v1 v2 = Some (v, t)" by (auto split:if_split_asm option.split_asm) - ultimately show "frame bal st'" by (auto split:if_split_asm) - qed - qed -next - case (25 e1 e2 e\<^sub>p e cd st) - show ?case (is "_ \ ?RHS") - proof - assume asm: "address e \ (STR ''Victim'') \ fmlookup e\<^sub>p (STR ''Victim'') = Some (victim, SKIP)" - show ?RHS - proof (rule allI[OF allI[OF allI[OF impI]]]) - fix rv1 and st' and bal - assume *: "frame bal st \ expr (AND e1 e2) e\<^sub>p e cd st = Normal (rv1, st')" - moreover from * obtain v1 t1 st'' where **: "expr e1 e\<^sub>p e cd (st\gas:=gas st - (costs\<^sub>e (AND e1 e2) e\<^sub>p e cd st)\) = Normal ((KValue v1, Value t1), st'')" - by (auto split:if_split_asm result.split_asm prod.split_asm Stackvalue.split_asm Type.split_asm) - moreover from * ** have ***: "frame bal st''" using 25(1) asm by (auto simp add:frame_def split:if_split_asm) - moreover from * ** *** obtain v2 t2 st''' where ****: "expr e2 e\<^sub>p e cd st'' = Normal ((KValue v2, Value t2), st''')" - by (auto split:if_split_asm result.split_asm prod.split_asm Stackvalue.split_asm Type.split_asm) - moreover from * ** *** **** have "frame bal st'''" using 25(2)[OF _ **] asm by (auto split:if_split_asm) - moreover from * ** **** obtain v t where "Valuetypes.vtand t1 t2 v1 v2 = Some (v, t)" by (auto split:if_split_asm option.split_asm) - ultimately show "frame bal st'" by (auto split:if_split_asm) - qed - qed -next - case (26 e1 e2 e\<^sub>p e cd st) - show ?case (is "_ \ ?RHS") - proof - assume asm: "address e \ (STR ''Victim'') \ fmlookup e\<^sub>p (STR ''Victim'') = Some (victim, SKIP)" - show ?RHS - proof (rule allI[OF allI[OF allI[OF impI]]]) - fix rv1 and st' and bal - assume *: "frame bal st \ expr (OR e1 e2) e\<^sub>p e cd st = Normal (rv1, st')" - moreover from * obtain v1 t1 st'' where **: "expr e1 e\<^sub>p e cd (st\gas:=gas st - (costs\<^sub>e (OR e1 e2) e\<^sub>p e cd st)\) = Normal ((KValue v1, Value t1), st'')" - by (auto split:if_split_asm result.split_asm prod.split_asm Stackvalue.split_asm Type.split_asm) - moreover from * ** have ***: "frame bal st''" using 26(1) asm by (auto simp add:frame_def split:if_split_asm) - moreover from * ** *** obtain v2 t2 st''' where ****: "expr e2 e\<^sub>p e cd st'' = Normal ((KValue v2, Value t2), st''')" - by (auto split:if_split_asm result.split_asm prod.split_asm Stackvalue.split_asm Type.split_asm) - moreover from * ** *** **** have "frame bal st'''" using 26(2)[OF _ **] asm by (auto split:if_split_asm) - moreover from * ** **** obtain v t where "Valuetypes.vtor t1 t2 v1 v2 = Some (v, t)" by (auto split:if_split_asm option.split_asm) - ultimately show "frame bal st'" by (auto split:if_split_asm) - qed - qed -next - case (27 i e\<^sub>p e cd st) - show ?case using 27(1)[of "()" "st\gas:=gas st - (costs\<^sub>e (LVAL i) e\<^sub>p e cd st)\"] apply safe by (auto simp add:frame_def split:if_split_asm) -next - case (28 i xe e\<^sub>p e cd st) - show ?case (is "_ \ ?RHS") - proof - assume asm: "address e \ (STR ''Victim'') \ fmlookup e\<^sub>p (STR ''Victim'') = Some (victim, SKIP)" - show ?RHS - proof (rule allI[OF allI[OF allI[OF impI]]]) - fix rv1 and st' and bal - assume *: "frame bal st \ expr (CALL i xe) e\<^sub>p e cd st = Normal (rv1, st')" - moreover from * have a1: "(applyf (costs\<^sub>e (CALL i xe) e\<^sub>p e cd) \ (\g. assert Gas (\st. gas st \ g) (modify (\st. st\gas := gas st - g\)))) st = Normal ((), st\gas := gas st - costs\<^sub>e (CALL i xe) e\<^sub>p e cd st\)" by auto - moreover from * obtain ct bla where **: "fmlookup e\<^sub>p (address e) = Some (ct, bla)" - by (auto split:if_split_asm option.split_asm) - moreover from * ** obtain fp f x where ***: "fmlookup ct i = Some (Method (fp, f, Some x))" - by (auto split:if_split_asm option.split_asm Member.split_asm) - moreover define e' where "e' = ffold_init ct (emptyEnv (address e) (sender e) (svalue e)) (fmdom ct)" - moreover from * ** *** obtain e'' cd' st'' st''' where ****: "load False fp xe e\<^sub>p e' emptyStore (st\gas:=gas st - (costs\<^sub>e (CALL i xe) e\<^sub>p e cd st), stack:=emptyStore\) e cd (st\gas:=gas st - (costs\<^sub>e (CALL i xe) e\<^sub>p e cd st)\) = Normal ((e'', cd', st''), st''')" - using e'_def by (auto split:if_split_asm result.split_asm) - moreover from * **** have f1: "frame bal st''" and ad: "address e' = address e''" - using asm 28(1)[OF _ ** _ *** _ _ _ _ e'_def, of _ "st\gas := gas st - costs\<^sub>e (CALL i xe) e\<^sub>p e cd st\" bla "(fp, f, Some x)" fp "(f, Some x)" f "Some x" x "st\gas := gas st - costs\<^sub>e (CALL i xe) e\<^sub>p e cd st, stack := emptyStore\" "st\gas := gas st - costs\<^sub>e (CALL i xe) e\<^sub>p e cd st\"] by (auto simp add:frame_def split:if_split_asm result.split_asm) - moreover from e'_def have ad2: "address e = address e'" using ffold_init_ad_same[of ct "(emptyEnv (address e) (sender e) (svalue e))" "(fmdom ct)" e'] by simp - moreover from * ** *** **** e'_def obtain st'''' where *****: "stmt f e\<^sub>p e'' cd' st'' = Normal ((), st'''')" by (auto split:if_split_asm result.split_asm) - moreover from f1 ad ad2 asm ***** have f2:"frame bal st''''" - using 28(2)[OF a1 ** _ *** _ _ _ _ e'_def _ ****, of bla "(fp, f, Some x)" "(f, Some x)" f "Some x" x e'' "(cd', st'')" "cd'" "st''" st''' st''' "()" st''] by (simp add:frame_def) - moreover from * ** *** **** ***** f1 f2 e'_def obtain rv st''''' where ******: "expr x e\<^sub>p e'' cd' st'''' = Normal (rv, st''''')" by (auto split:if_split_asm result.split_asm) - ultimately have "st' = st'''''\stack:=stack st''', memory := memory st'''\" apply safe by auto - moreover from f1 f2 ad ad2 asm a1 ***** ****** have "\rv4 st4' bal. - frame bal st'''' \ - local.expr x e\<^sub>p e'' cd' st'''' = Normal (rv4, st4') \ - frame bal st4'" using e'_def asm 28(3)[OF a1 ** _ *** _ _ _ _ e'_def _ **** _ _ _ _ *****, of bla "(fp, f, Some x)" " (f, Some x)" "Some x" x "(cd', st'')" st'' st''' st''' "()"] apply safe by auto - with ****** f2 have "frame bal st'''''" by blast - ultimately show "frame bal st'" by (simp add:frame_def) - qed - qed -next - case (29 ad i xe val e\<^sub>p e cd st) - show ?case (is "_ \ ?RHS") - proof - assume asm: "address e \ (STR ''Victim'') \ fmlookup e\<^sub>p (STR ''Victim'') = Some (victim, SKIP)" - show ?RHS - proof (rule allI[OF allI[OF allI[OF impI]]]) - fix rv1 and st' and bal - assume *: "frame bal st \ expr (ECALL ad i xe val) e\<^sub>p e cd st = Normal (rv1, st')" - moreover from * have a1: "(applyf (costs\<^sub>e (ECALL ad i xe val) e\<^sub>p e cd) \ (\g. assert Gas (\st. gas st \ g) (modify (\st. st\gas := gas st - g\)))) st = Normal ((), st\gas := gas st - costs\<^sub>e (ECALL ad i xe val) e\<^sub>p e cd st\)" by auto - moreover from * obtain adv st'' where **: "expr ad e\<^sub>p e cd (st\gas:=gas st - (costs\<^sub>e (ECALL ad i xe val) e\<^sub>p e cd st)\) = Normal ((KValue adv, Value TAddr), st'')" - by (auto split:if_split_asm result.split_asm Stackvalue.split_asm Type.split_asm Types.split_asm) - moreover from * ** have f1: "frame bal st''"using asm 29(1) by (auto simp add:frame_def split:if_split_asm) - moreover from * ** obtain ct bla where ***: "fmlookup e\<^sub>p adv = Some (ct, bla)" - by (auto split:if_split_asm option.split_asm) - moreover from * ** *** obtain fp f x where ****: "fmlookup ct i = Some (Method (fp, f, Some x))" - by (auto split:if_split_asm option.split_asm Member.split_asm) - moreover from * ** *** **** obtain v t st''' where *****: "expr val e\<^sub>p e cd st'' = Normal ((KValue v, Value t), st''')" by (auto split:if_split_asm result.split_asm Stackvalue.split_asm Type.split_asm) - moreover from f1 ***** asm have f2: "frame bal st'''" and f3: "frame bal (st'''\stack := emptyStore, memory := emptyStore\)" using asm 29(2)[OF a1 ** _ _ _ _ *** _ ****] by (auto simp add:frame_def) - moreover define e' where "e' = ffold_init ct (emptyEnv adv (address e) v) (fmdom ct)" - moreover from * ** *** **** ***** obtain e'' cd' st'''' st''''' where *******: "load True fp xe e\<^sub>p e' emptyStore (st'''\stack:=emptyStore, memory:=emptyStore\) e cd st''' = Normal ((e'', cd', st''''), st''''')" - using e'_def by (auto split:if_split_asm result.split_asm option.split_asm) - moreover have "(\ev cda st st' bal. - local.load True fp xe e\<^sub>p e' emptyStore (st'''\stack := emptyStore, memory := emptyStore\) e cd st''' = Normal ((ev, cda, st), st') \ - (frame bal (st'''\stack := emptyStore, memory := emptyStore\) \ frame bal st) \ - (frame bal st''' \ frame bal st') \ address e' = address ev \ sender e' = sender ev \ svalue e' = svalue ev)" - using 29(3)[OF a1 ** _ _ _ _ *** _ **** _ _ _ _ ***** _ _ _ e'_def, of "KValue adv" "Value TAddr" TAddr bla "(fp, f, Some x)" fp "(f, Some x)" f "Some x" x "KValue v" "Value t" t "st'''\stack := emptyStore, memory := emptyStore\" st'''] asm ******* by simp - then have "frame bal st'''' \ frame bal st''''' \ address e' = address e''" using ******* f2 f3 by blast - then have f4: "frame bal st''''" and ad1: "address e' = address e''" by auto - moreover from * ** *** **** ***** ******* e'_def obtain acc where ******: "Accounts.transfer (address e) adv v (accounts st'''') = Some acc" by (auto split:if_split_asm option.split_asm) - then have ******: "Accounts.transfer (address e) adv v (accounts st'''') = Some acc" by (auto split:if_split_asm option.split_asm) - moreover from f4 have f5: "frame bal (st''''\accounts := acc\)" using transfer_frame[OF ******] asm by simp - moreover from e'_def have ad2: "adv = address e'" using ffold_init_ad_same[of ct "(emptyEnv adv (address e) v)" "(fmdom ct)" e'] by simp - moreover from * ** *** **** ***** ****** ******* ****** obtain st'''''' where ********: "stmt f e\<^sub>p e'' cd' (st''''\accounts := acc\) = Normal ((), st'''''')" - using e'_def by (auto simp del: transfer.simps split:if_split_asm result.split_asm) - moreover have "adv \ STR ''Victim''" - proof (rule ccontr) - assume "\ adv \ STR ''Victim''" - with asm ** *** **** show False using victim_def fmap_of_list_SomeD[of "[(STR ''balance'', Var (STMap TAddr (STValue (TUInt 256)))), (STR ''deposit'', Method ([], deposit, None)), (STR ''withdraw'', Method ([], keep, None))]"] by auto - qed - with ad1 ad2 have ad: "address e'' \ STR ''Victim'' \ fmlookup e\<^sub>p (STR ''Victim'') = Some (victim, SKIP)" using asm by simp - then have "(\bal. frame bal (st''''\accounts := acc\) \ frame bal st'''''')" using 29(4)[OF a1 ** _ _ _ _ *** _ **** _ _ _ _ ***** _ _ _ _ _ ******* _ _ _ ******, of "KValue adv" "Value TAddr" TAddr bla "(fp, f, Some x)" "(f, Some x)" f "Some x" x "KValue v" "Value t" t e'' "(cd', st'''')" cd' st''''' st''''' "()" "st''''\accounts := acc\"] ******** e'_def by auto - then have f4: "frame bal st''''''" using f5 ******** by auto - moreover from * ** *** **** ***** ****** ******* ******** obtain rv st''''''' where *********: "expr x e\<^sub>p e'' cd' st'''''' = Normal (rv, st''''''')" - using e'_def by (auto split:if_split_asm result.split_asm) - ultimately have "st' = st'''''''\stack:=stack st''''', memory := memory st'''''\" apply safe by auto - moreover from ad have "\rv4 st4' bal. - frame bal st'''''' \ - local.expr x e\<^sub>p e'' cd' st'''''' = Normal (rv4, st4') \ - frame bal st4'" - using e'_def 29(5)[OF a1 ** _ _ _ _ *** _ **** _ _ _ _ ***** _ _ _ _ _ ******* _ _ _ ****** _ ********, of "KValue adv" "Value TAddr" TAddr bla "(fp, f, Some x)"] by auto - then have"frame bal st'''''''" using f4 ********* by blast - ultimately show "frame bal st'" by (simp add:frame_def) - qed - qed -next - case (30 cp i\<^sub>p t\<^sub>p pl e el e\<^sub>p e\<^sub>v' cd' st' e\<^sub>v cd st) - show ?case (is "_ \ ?RHS") - proof - assume asm: "address e\<^sub>v \ (STR ''Victim'') \ fmlookup e\<^sub>p (STR ''Victim'') = Some (victim, SKIP)" - show ?RHS - proof (rule allI[OF allI[OF allI[OF allI[OF allI[OF impI]]]]]) - fix ev and cda and sta and st'a and bal - assume *: "local.load cp ((i\<^sub>p, t\<^sub>p) # pl) (e # el) e\<^sub>p e\<^sub>v' cd' st' e\<^sub>v cd st = Normal ((ev, cda, sta), st'a)" - moreover from * obtain v t st'' where **: "expr e e\<^sub>p e\<^sub>v cd st = Normal ((v,t),st'')" by (auto split: result.split_asm) - moreover from * ** obtain cd'' e\<^sub>v'' st''' where ***: "decl i\<^sub>p t\<^sub>p (Some (v,t)) cp cd (memory st'') cd' e\<^sub>v' st' = Normal ((cd'', e\<^sub>v''),st''')" by (auto split: result.split_asm) - moreover from *** have ad: "address e\<^sub>v' = address e\<^sub>v'' \ sender e\<^sub>v' = sender e\<^sub>v'' \ svalue e\<^sub>v' = svalue e\<^sub>v''" using decl_gas_address by simp - moreover from * ** *** obtain ev' cda' sta' st'a' where ****: "local.load cp pl el e\<^sub>p e\<^sub>v'' cd'' st''' e\<^sub>v cd st''= Normal ((ev', cda', sta'), st'a')" by (auto split: result.split_asm) - ultimately have "ev = ev'" and "sta = sta'" and "st'a = st'a'" by simp+ - - from **** asm have IH: "(frame bal st''' \ frame bal sta') \ - (frame bal st'' \ frame bal st'a') \ - address e\<^sub>v'' = address ev' \ sender e\<^sub>v'' = sender ev' \ svalue e\<^sub>v'' = svalue ev'" using 30(2)[OF ** _ _ _ ***, of st'' "()" cd'' e\<^sub>v'' st''' st''' "()" st''] apply safe by (auto simp add:frame_def) - show "(frame bal st' \ frame bal sta) \ (frame bal st \ frame bal st'a) \ address e\<^sub>v' = address ev \ sender e\<^sub>v' = sender ev \ svalue e\<^sub>v' = svalue ev" - proof (rule conj3) - show "frame bal st' \ frame bal sta" - proof - assume "frame bal st'" - with * ** *** have "frame bal st'''" using decl_frame by simp - with IH have "frame bal sta'" by simp - with `sta = sta'` show "frame bal sta" by simp - qed - next - show "frame bal st \ frame bal st'a" - proof - assume "frame bal st" - with ** have "frame bal st''" using 30(1) asm by simp - with IH have "frame bal st'a'" by simp - with `st'a = st'a'` show "frame bal st'a" by simp - qed - next - from ad IH show "address e\<^sub>v' = address ev \ sender e\<^sub>v' = sender ev \ svalue e\<^sub>v' = svalue ev" using `ev = ev'` by simp - qed - qed - qed -next - case (31 vv vw vx vy vz wa wb wc wd st) - then show ?case by simp -next - case (32 we wf wg wh wi wj wk wl wm st) - then show ?case by simp -next - case (33 wn wo e\<^sub>v' cd' st' e\<^sub>v cd st) - then show ?case by simp -next - case (34 i e\<^sub>p e cd st) - show ?case (is "_ \ ?RHS") - proof - assume asm: "address e \ (STR ''Victim'') \ fmlookup e\<^sub>p (STR ''Victim'') = Some (victim, SKIP)" - show ?RHS - proof (rule allI[OF allI[OF allI[OF impI]]]) - fix rv3 and st3' and bal - assume *: "frame bal st \ local.rexp (L.Id i) e\<^sub>p e cd st = Normal (rv3, st3')" - show "frame bal st3'" - proof (cases "fmlookup (denvalue e) i") - case None - with * show ?thesis by simp - next - case (Some a) - then show ?thesis - proof (cases a) - case (Pair tp b) - then show ?thesis - proof (cases b) - case (Stackloc l) - then show ?thesis - proof (cases "accessStore l (stack st)") - case None - with * Some Pair Stackloc show ?thesis by (auto split: Type.split_asm STypes.split_asm) - next - case s2: (Some a) - with * Some Pair Stackloc s2 show ?thesis by (auto split: Type.split_asm STypes.split_asm Stackvalue.split_asm) - qed - next - case (Storeloc x2) - with * Some Pair Storeloc show ?thesis by (auto split: Type.split_asm STypes.split_asm option.split_asm) - qed - qed - qed - qed - qed -next - case (35 i r e\<^sub>p e cd st) - show ?case (is "_ \ ?RHS") - proof - assume asm: "address e \ (STR ''Victim'') \ fmlookup e\<^sub>p (STR ''Victim'') = Some (victim, SKIP)" - show ?RHS - proof (rule allI[OF allI[OF allI[OF impI]]]) - fix rv3 and st3' and bal - assume *: "frame bal st \ local.rexp (L.Ref i r) e\<^sub>p e cd st = Normal (rv3, st3')" - show "frame bal st3'" - proof (cases "fmlookup (denvalue e) i") - case None - with * show ?thesis by simp - next - case (Some a) - then show ?thesis - proof (cases a) - case (Pair tp b) - then show ?thesis - proof (cases b) - case (Stackloc l') - then show ?thesis - proof (cases "accessStore l' (stack st)") - case None - with * Some Pair Stackloc show ?thesis by simp - next - case s2: (Some a) - then show ?thesis - proof (cases a) - case (KValue x1) - with * Some Pair Stackloc s2 show ?thesis by simp - next - case (KCDptr l'') - then show ?thesis - proof (cases tp) - case (Value x1) - with * Some Pair Stackloc s2 KCDptr show ?thesis by simp - next - case (Calldata t) - with * Some Pair Stackloc s2 KCDptr obtain l''' t' st' where **: "msel False t l'' r e\<^sub>p e cd st = Normal ((l''',t'), st')" by (auto split: Type.split_asm STypes.split_asm result.split_asm) - then have "\rv1 st1' bal. - frame bal st \ - local.msel False t l'' r e\<^sub>p e cd st = Normal (rv1, st1') \ - frame bal st1'" using asm 35(1)[OF Some Pair Stackloc _ s2 KCDptr Calldata] by auto - with * ** have f2: "frame bal st'" by blast - then show ?thesis - proof (cases t') - case (MTArray x t'') - then show ?thesis - proof (cases "accessStore l''' cd") - case None - with * ** Some Pair Stackloc s2 KCDptr Calldata MTArray show ?thesis by simp - next - case s3: (Some a) - then show ?thesis - proof (cases a) - case (MValue x1) - with * ** Some Pair Stackloc s2 KCDptr Calldata MTArray s3 show ?thesis by simp - next - case (MPointer x2) - with * ** f2 Some Pair Stackloc s2 KCDptr Calldata MTArray s3 show ?thesis by simp - qed - qed - next - case (MTValue t''') - then show ?thesis - proof (cases "accessStore l''' cd") - case None - with * ** Some Pair Stackloc s2 KCDptr Calldata MTValue show ?thesis by simp - next - case s3: (Some a) - then show ?thesis - proof (cases a) - case (MValue x1) - with * ** f2 Some Pair Stackloc s2 KCDptr Calldata MTValue s3 show ?thesis by simp - next - case (MPointer x2) - with * ** Some Pair Stackloc s2 KCDptr Calldata MTValue s3 show ?thesis by simp - qed - qed - qed - next - case (Memory x3) - with * Some Pair Stackloc s2 KCDptr show ?thesis by simp - next - case (Storage x4) - with * Some Pair Stackloc s2 KCDptr show ?thesis by simp - qed - next - case (KMemptr l'') - then show ?thesis - proof (cases tp) - case (Value x1) - with * Some Pair Stackloc s2 KMemptr show ?thesis by simp - next - case (Calldata x2) - with * Some Pair Stackloc s2 KMemptr show ?thesis by simp - next - case (Memory t) - with * Some Pair Stackloc s2 KMemptr obtain l''' t' st' where **: "msel True t l'' r e\<^sub>p e cd st = Normal ((l''',t'), st')" by (auto split: Type.split_asm STypes.split_asm result.split_asm) - then have "\rv1 st1' bal. frame bal st \ - local.msel True t l'' r e\<^sub>p e cd st = Normal (rv1, st1') \ - frame bal st1'" using asm 35(2)[OF Some Pair Stackloc _ s2 KMemptr Memory, of st] by auto - with * ** have f2: "frame bal st'" by blast - then show ?thesis - proof (cases t') - case (MTArray x11 x12) - then show ?thesis - proof (cases "accessStore l''' (memory st')") - case None - with * ** Some Pair Stackloc s2 KMemptr Memory MTArray show ?thesis by simp - next - case s3: (Some a) - then show ?thesis - proof (cases a) - case (MValue x1) - with * ** Some Pair Stackloc s2 KMemptr Memory MTArray s3 show ?thesis by simp - next - case (MPointer x2) - with * ** f2 Some Pair Stackloc s2 KMemptr Memory MTArray s3 show ?thesis by simp - qed - qed - next - case (MTValue x2) - then show ?thesis - proof (cases "accessStore l''' (memory st')") - case None - with * ** Some Pair Stackloc s2 KMemptr Memory MTValue show ?thesis by simp - next - case s3: (Some a) - then show ?thesis - proof (cases a) - case (MValue x1) - with * ** f2 Some Pair Stackloc s2 KMemptr Memory MTValue s3 show ?thesis by simp - next - case (MPointer x2) - with * ** Some Pair Stackloc s2 KMemptr Memory MTValue s3 show ?thesis by simp - qed - qed - qed - next - case (Storage x4) - with * Some Pair Stackloc s2 KMemptr show ?thesis by simp - qed - next - case (KStoptr l'') - then show ?thesis - proof (cases tp) - case (Value x1) - with * Some Pair Stackloc s2 KStoptr show ?thesis by simp - next - case (Calldata x2) - with * Some Pair Stackloc s2 KStoptr show ?thesis by simp - next - case (Memory x3) - with * Some Pair Stackloc s2 KStoptr show ?thesis by simp - next - case (Storage t) - with * Some Pair Stackloc s2 KStoptr obtain l''' t' st' where **: "ssel t l'' r e\<^sub>p e cd st = Normal ((l''',t'), st')" by (auto split: Type.split_asm STypes.split_asm result.split_asm) - then have "\rv2 st2' bal. - frame bal st \ - local.ssel t l'' r e\<^sub>p e cd st = Normal (rv2, st2') \ - frame bal st2'" using asm 35(3)[OF Some Pair Stackloc _ s2 KStoptr Storage, of st] by auto - with * ** have "frame bal st'" by blast - with * ** Some Pair Stackloc s2 KStoptr Storage show ?thesis by (simp split: STypes.split_asm option.split_asm) - qed - qed - qed - next - case (Storeloc l') - then show ?thesis - proof (cases tp) - case (Value x1) - with * Some Pair Storeloc show ?thesis by simp - next - case (Calldata x2) - with * Some Pair Storeloc show ?thesis by simp - next - case (Memory x3) - with * Some Pair Storeloc show ?thesis by simp - next - case (Storage t) - with * Some Pair Storeloc obtain l'' t' st' where **: "ssel t l' r e\<^sub>p e cd st = Normal ((l'',t'), st')" by (auto split: result.split_asm) - then have "\rv2 st2' bal. - frame bal st \ - local.ssel t l' r e\<^sub>p e cd st = Normal (rv2, st2') \ - frame bal st2'" using asm 35(4)[OF Some Pair Storeloc Storage] by auto - with * ** have "frame bal st'" by blast - with * ** Some Pair Storeloc Storage show ?thesis by (simp split: STypes.split_asm option.split_asm) - qed - qed - qed - qed - qed - qed -next - case (36 e\<^sub>p e cd st) - show ?case (is "?LHS \ ?RHS") - proof - assume *: "fmlookup e\<^sub>p STR ''Victim'' = Some (victim, SKIP)" - show ?RHS (is "\st6'. ?RHS st6'") - proof - fix st6' - show "?RHS st6'" (is "?LHS \ ?RHS") - proof - assume t0: "stmt SKIP e\<^sub>p e cd st = Normal ((), st6')" - show ?RHS (is "?LHS \ ?RHS") - proof - show "?LHS" - proof - assume ad: "address e \ STR ''Victim''" - show "\bal. frame bal st \ frame bal st6'" - proof - fix bal - show "frame bal st \ frame bal st6'" - proof - assume "frame bal st" - with t0 * show "frame bal st6'" by (auto simp add: frame_def split:if_split_asm) - qed - qed - qed - next - show "?RHS" (is "?LHS \ ?RHS") - proof - assume "address e = STR ''Victim''" - show ?RHS (is "?A \ (?B \ ?C)") - proof (rule conj3) - show ?A (is "\s val bal x. ?LHS s val bal x") - proof (rule allI[OF allI[OF allI[OF allI]]]) - fix s val bal x - show "?LHS s val bal x" (is "?LHS \ ?RHS") - proof - assume ?LHS - then show ?RHS by simp - qed - qed - next - show ?B (is "\s bal x. ?LHS s bal x") - proof (rule allI[OF allI[OF allI]]) - fix s bal x - show "?LHS s bal x" (is "?LHS \ ?RHS") - proof - assume ?LHS - then show ?RHS by simp - qed - qed - next - show ?C (is "\s bal. ?LHS s bal") - proof (rule allI[OF allI]) - fix s bal - show "?LHS s bal" (is "?LHS \ ?RHS") - proof - assume ?LHS - then show ?RHS by simp - qed - qed - qed - qed - qed - qed - qed - qed -next - case (37 lv ex e\<^sub>p env cd st) - show ?case (is "?LHS \ ?RHS") - proof - assume 0: "fmlookup e\<^sub>p STR ''Victim'' = Some (victim, SKIP)" - show ?RHS (is "\st6'. ?RHS st6'") - proof - fix st6' - show "?RHS st6'" (is "?LHS \ ?RHS") - proof - assume *: "stmt (ASSIGN lv ex) e\<^sub>p env cd st = Normal ((), st6')" - show ?RHS (is "?LHS \ ?RHS") - proof - show "?LHS" - proof - assume asm: "address env \ STR ''Victim''" - show "\bal. frame bal st \ frame bal st6'" - proof - fix bal - show "frame bal st \ frame bal st6'" - proof - assume "frame bal st" - with * have a1: "(applyf (costs (ASSIGN lv ex) e\<^sub>p env cd) \ (\g. assert Gas (\st. gas st \ g) (modify (\st. st\gas := gas st - g\)))) st = - Normal ((), st\gas:=gas st - costs (ASSIGN lv ex) e\<^sub>p env cd st\)" - and f1: "frame bal (st\gas:=gas st - costs (ASSIGN lv ex) e\<^sub>p env cd st\)" by (auto simp add:frame_def) - moreover from * obtain kv kt st' where **: "expr ex e\<^sub>p env cd (st\gas:=gas st - costs (ASSIGN lv ex) e\<^sub>p env cd st\) = Normal ((kv, kt), st')" by (auto split:if_split_asm result.split_asm) - ultimately have "\rv4 st4' (ev4'::Environment) bal. - frame bal (st\gas := gas st - costs (ASSIGN lv ex) e\<^sub>p env cd st\) \ - local.expr ex e\<^sub>p env cd (st\gas := gas st - costs (ASSIGN lv ex) e\<^sub>p env cd st\) = Normal (rv4, st4') \ - frame bal st4'" using asm 0 37(1) by simp - with f1 ** have f2: "frame bal st'" by blast - show "frame bal st6'" - proof (cases kv) - case (KValue v) - then show ?thesis - proof (cases kt) - case (Value t) - with * ** KValue obtain rv rt st'' where ***: "lexp lv e\<^sub>p env cd st' = Normal ((rv,rt), st'')" by (auto split:if_split_asm result.split_asm) - with KValue Value have "\rv5 st5' (ev5'::Environment) bal. - frame bal st' \ - local.lexp lv e\<^sub>p env cd st' = Normal (rv5, st5') \ - frame bal st5'" using asm 0 37(2)[OF a1 **] by simp - with f2 *** have f3: "frame bal st''" by blast - then show ?thesis - proof (cases rv) - case (LStackloc l') - then show ?thesis - proof (cases rt) - case v2: (Value t') - then show ?thesis - proof (cases "Valuetypes.convert t t' v") - case None - with * ** *** KValue Value LStackloc v2 show ?thesis by (auto split:if_split_asm) - next - case (Some a) - then show ?thesis - proof (cases a) - case (Pair v' b) - with * ** *** KValue Value LStackloc v2 Some have "st6' = st'' \stack := updateStore l' (KValue v') (stack st'')\" by (auto split:if_split_asm) - with f3 show ?thesis by (simp add:frame_def) - qed - qed - next - case (Calldata x2) - with * ** *** KValue Value LStackloc show ?thesis by (auto split:if_split_asm) - next - case (Memory x3) - with * ** *** KValue Value LStackloc show ?thesis by (auto split:if_split_asm) - next - case (Storage x4) - with * ** *** KValue Value LStackloc show ?thesis by (auto split:if_split_asm) - qed - next - case (LMemloc l') - then show ?thesis - proof (cases rt) - case v2: (Value t') - with * ** *** KValue Value LMemloc show ?thesis by (auto split:if_split_asm) - next - case (Calldata x2) - with * ** *** KValue Value LMemloc show ?thesis by (auto split:if_split_asm) - next - case (Memory x3) - then show ?thesis - proof (cases x3) - case (MTArray x11 x12) - with * ** *** KValue Value LMemloc Memory show ?thesis by (auto split:if_split_asm) - next - case (MTValue t') - then show ?thesis - proof (cases "Valuetypes.convert t t' v") - case None - with * ** *** KValue Value LMemloc Memory MTValue show ?thesis by (auto split:if_split_asm) - next - case (Some a) - then show ?thesis - proof (cases a) - case (Pair v' b) - with * ** *** KValue Value LMemloc Memory MTValue Some have "st6' = st'' \memory := updateStore l' (MValue v') (memory st'')\" by (auto split:if_split_asm) - with f3 show ?thesis by (simp add:frame_def) - qed - qed - qed - next - case (Storage x4) - with * ** *** KValue Value LMemloc Storage show ?thesis by (auto split:if_split_asm) - qed - next - case (LStoreloc l') - then show ?thesis - proof (cases rt) - case v2: (Value x1) - with * ** *** KValue Value LStoreloc show ?thesis by (auto split:if_split_asm) - next - case (Calldata x2) - with * ** *** KValue Value LStoreloc show ?thesis by (auto split:if_split_asm) - next - case (Memory x3) - with * ** *** KValue Value LStoreloc show ?thesis by (auto split:if_split_asm) - next - case (Storage x4) - then show ?thesis - proof (cases x4) - case (STArray x11 x12) - with * ** *** KValue Value LStoreloc Storage show ?thesis by (auto split:if_split_asm) - next - case (STMap x21 x22) - with * ** *** KValue Value LStoreloc Storage show ?thesis by (auto split:if_split_asm) - next - case (STValue t') - then show ?thesis - proof (cases "Valuetypes.convert t t' v") - case None - with * ** *** KValue Value LStoreloc Storage STValue show ?thesis by (auto split:if_split_asm) - next - case (Some a) - then show ?thesis - proof (cases a) - case (Pair v' b) - then show ?thesis - proof (cases "fmlookup (storage st'') (address env)") - case None - with * ** *** KValue Value LStoreloc Storage STValue Some Pair show ?thesis by (auto split:if_split_asm) - next - case s2: (Some s) - with * ** *** KValue Value LStoreloc Storage STValue Some Pair have "st6' = st''\storage := fmupd (address env) (fmupd l' v' s) (storage st'')\" by (auto split:if_split_asm) - with f3 show ?thesis using asm by (simp add:frame_def) - qed - qed - qed - qed - qed - qed - next - case (Calldata x2) - with * ** KValue show ?thesis by (auto split:if_split_asm) - next - case (Memory x3) - with * ** KValue show ?thesis by (auto split:if_split_asm) - next - case (Storage x4) - with * ** KValue show ?thesis by (auto split:if_split_asm) - qed - next - case (KCDptr p) - then show ?thesis - proof (cases kt) - case (Value t) - with * ** KCDptr show ?thesis by (auto split:if_split_asm) - next - case (Calldata x2) - then show ?thesis - proof (cases x2) - case (MTArray x t) - with * ** KCDptr Calldata obtain rv rt st'' where ***: "lexp lv e\<^sub>p env cd st' = Normal ((rv,rt), st'')" by (auto split:if_split_asm result.split_asm) - with KCDptr Calldata MTArray have "\rv5 st5' (ev5'::Environment) bal. - frame bal st' \ - local.lexp lv e\<^sub>p env cd st' = Normal (rv5, st5') \ - frame bal st5'" using asm 0 37(3)[OF a1 **] by auto - with f2 *** have f3: "frame bal st''" by blast - then show ?thesis - proof (cases rv) - case (LStackloc l') - then show ?thesis - proof (cases rt) - case (Value x1) - with * ** *** KCDptr Calldata MTArray LStackloc show ?thesis by (auto split:if_split_asm) - next - case c2: (Calldata x2) - with * ** *** KCDptr Calldata MTArray LStackloc show ?thesis by (auto split:if_split_asm) - next - case (Memory x3) - with f3 * ** *** KCDptr Calldata MTArray LStackloc show ?thesis by (auto simp add:frame_def split:if_split_asm) - next - case (Storage x4) - then show ?thesis - proof (cases "accessStore l' (stack st'')") - case None - with * ** *** KCDptr Calldata MTArray LStackloc Storage show ?thesis by (simp split:if_split_asm) - next - case (Some sv) - then show ?thesis - proof (cases sv) - case (KValue x1) - with * ** *** KCDptr Calldata MTArray LStackloc Storage Some show ?thesis by (simp split:if_split_asm) - next - case c2: (KCDptr x2) - with * ** *** KCDptr Calldata MTArray LStackloc Storage Some show ?thesis by (simp split:if_split_asm) - next - case (KMemptr x3) - with * ** *** KCDptr Calldata MTArray LStackloc Storage Some show ?thesis by (simp split:if_split_asm) - next - case (KStoptr p') - then show ?thesis - proof (cases "fmlookup (storage st'') (address env)") - case None - with * ** *** KCDptr Calldata MTArray LStackloc Storage Some KStoptr show ?thesis by (simp split:if_split_asm) - next - case s2: (Some s) - then show ?thesis - proof (cases "cpm2s p p' x t cd s") - case None - with * ** *** KCDptr Calldata MTArray LStackloc Storage Some KStoptr s2 show ?thesis by (simp split:if_split_asm) - next - case s3: (Some s') - with * ** *** KCDptr Calldata MTArray LStackloc Storage Some KStoptr s2 have "st6' = st'' \storage := fmupd (address env) s' (storage st'')\" by (auto split:if_split_asm) - with f3 show ?thesis using asm by (simp add:frame_def) - qed - qed - qed - qed - qed - next - case (LMemloc l') - then show ?thesis - proof (cases "cpm2m p l' x t cd (memory st'')") - case None - with * ** *** KCDptr Calldata MTArray LMemloc show ?thesis by (auto split:if_split_asm) - next - case (Some m) - with * ** *** KCDptr Calldata MTArray LMemloc have "st6' = st'' \memory := m\" by (auto split:if_split_asm) - with f3 show ?thesis using asm by (simp add:frame_def) - qed - next - case (LStoreloc l') - then show ?thesis - proof (cases "fmlookup (storage st'') (address env)") - case None - with * ** *** KCDptr Calldata MTArray LStoreloc show ?thesis by (auto split:if_split_asm) - next - case (Some s) - then show ?thesis - proof (cases "cpm2s p l' x t cd s") - case None - with * ** *** KCDptr Calldata MTArray LStoreloc Some show ?thesis by (auto split:if_split_asm) - next - case s2: (Some s') - with * ** *** KCDptr Calldata MTArray LStoreloc Some s2 have "st6' = st'' \storage := fmupd (address env) s' (storage st'')\" by (auto split:if_split_asm) - with f3 show ?thesis using asm by (simp add:frame_def) - qed - qed - qed - next - case (MTValue x2) - with * ** KCDptr Calldata show ?thesis by (simp split:if_split_asm) - qed - next - case (Memory x3) - with * ** KCDptr show ?thesis by (simp split:if_split_asm) - next - case (Storage x4) - with * ** KCDptr show ?thesis by (simp split:if_split_asm) - qed - next - case (KMemptr p) - then show ?thesis - proof (cases kt) - case (Value t) - with * ** KMemptr show ?thesis by (auto split:if_split_asm) - next - case (Calldata x2) - with * ** KMemptr show ?thesis by (simp split:if_split_asm) - next - case (Memory x3) - then show ?thesis - proof (cases x3) - case (MTArray x t) - with * ** KMemptr Memory obtain rv rt st'' where ***: "lexp lv e\<^sub>p env cd st' = Normal ((rv,rt), st'')" by (auto split:if_split_asm result.split_asm) - with KMemptr Memory MTArray have "\rv5 st5' (ev5'::Environment) bal. - frame bal st' \ - local.lexp lv e\<^sub>p env cd st' = Normal (rv5, st5') \ - frame bal st5'" using asm 0 37(4)[OF a1 **] by auto - with f2 *** have f3: "frame bal st''" by blast - then show ?thesis - proof (cases rv) - case (LStackloc l') - then show ?thesis - proof (cases rt) - case (Value x1) - with * ** *** KMemptr Memory MTArray LStackloc show ?thesis by (auto split:if_split_asm) - next - case (Calldata x2) - with * ** *** KMemptr Memory MTArray LStackloc show ?thesis by (auto split:if_split_asm) - next - case m3: (Memory x3) - with f3 * ** *** KMemptr Memory MTArray LStackloc show ?thesis by (auto simp add:frame_def split:if_split_asm) - next - case (Storage x4) - then show ?thesis - proof (cases "accessStore l' (stack st'')") - case None - with * ** *** KMemptr Memory MTArray LStackloc Storage show ?thesis by (simp split:if_split_asm) - next - case (Some sv) - then show ?thesis - proof (cases sv) - case (KValue x1) - with * ** *** KMemptr Memory MTArray LStackloc Storage Some show ?thesis by (simp split:if_split_asm) - next - case (KCDptr x2) - with * ** *** KMemptr Memory MTArray LStackloc Storage Some show ?thesis by (simp split:if_split_asm) - next - case m2: (KMemptr x3) - with * ** *** KMemptr Memory MTArray LStackloc Storage Some show ?thesis by (simp split:if_split_asm) - next - case (KStoptr p') - then show ?thesis - proof (cases "fmlookup (storage st'') (address env)") - case None - with * ** *** KMemptr Memory MTArray LStackloc Storage Some KStoptr show ?thesis by (simp split:if_split_asm) - next - case s2: (Some s) - then show ?thesis - proof (cases "cpm2s p p' x t (memory st'') s") - case None - with * ** *** KMemptr Memory MTArray LStackloc Storage Some KStoptr s2 show ?thesis by (simp split:if_split_asm) - next - case s3: (Some s') - with * ** *** KMemptr Memory MTArray LStackloc Storage Some KStoptr s2 have "st6' = st'' \storage := fmupd (address env) s' (storage st'')\" by (auto split:if_split_asm) - with f3 show ?thesis using asm by (simp add:frame_def) - qed - qed - qed - qed - qed - next - case (LMemloc l') - with * ** *** KMemptr Memory MTArray LMemloc have "st6' = st'' \memory := updateStore l' (MPointer p) (memory st'')\" by (auto split:if_split_asm) - with f3 show ?thesis using asm by (simp add:frame_def) - next - case (LStoreloc l') - then show ?thesis - proof (cases "fmlookup (storage st'') (address env)") - case None - with * ** *** KMemptr Memory MTArray LStoreloc show ?thesis by (auto split:if_split_asm) - next - case (Some s) - then show ?thesis - proof (cases "cpm2s p l' x t (memory st'') s") - case None - with * ** *** KMemptr Memory MTArray LStoreloc Some show ?thesis by (auto split:if_split_asm) - next - case s2: (Some s') - with * ** *** KMemptr Memory MTArray LStoreloc Some s2 have "st6' = st'' \storage := fmupd (address env) s' (storage st'')\" by (auto split:if_split_asm) - with f3 show ?thesis using asm by (simp add:frame_def) - qed - qed - qed - next - case (MTValue x2) - with * ** KMemptr Memory show ?thesis by (simp split:if_split_asm) - qed - next - case (Storage x4) - with * ** KMemptr show ?thesis by (simp split:if_split_asm) - qed - next - case (KStoptr p) - then show ?thesis - proof (cases kt) - case (Value t) - with * ** KStoptr show ?thesis by (auto split:if_split_asm) - next - case (Calldata x2) - with * ** KStoptr show ?thesis by (simp split:if_split_asm) - next - case (Storage x3) - then show ?thesis - proof (cases x3) - case (STArray x t) - with * ** KStoptr Storage obtain rv rt st'' where ***: "lexp lv e\<^sub>p env cd st' = Normal ((rv,rt), st'')" by (auto split:if_split_asm result.split_asm) - with KStoptr Storage STArray have "\rv5 st5' bal. - frame bal st' \ - local.lexp lv e\<^sub>p env cd st' = Normal (rv5, st5') \ - frame bal st5'" using asm 0 37(5)[OF a1 **] by auto - with f2 *** have f3: "frame bal st''" by blast - then show ?thesis - proof (cases rv) - case (LStackloc l') - then show ?thesis - proof (cases rt) - case (Value x1) - with * ** *** KStoptr Storage STArray LStackloc show ?thesis by (auto split:if_split_asm) - next - case (Calldata x2) - with * ** *** KStoptr Storage STArray LStackloc show ?thesis by (auto split:if_split_asm) - next - case (Memory x3) - then show ?thesis - proof (cases "accessStore l' (stack st'')") - case None - with * ** *** KStoptr Storage STArray LStackloc Memory show ?thesis by (simp split:if_split_asm) - next - case (Some sv) - then show ?thesis - proof (cases sv) - case (KValue x1) - with * ** *** KStoptr Storage STArray LStackloc Memory Some show ?thesis by (simp split:if_split_asm) - next - case (KCDptr x2) - with * ** *** KStoptr Storage STArray LStackloc Memory Some show ?thesis by (simp split:if_split_asm) - next - case (KMemptr p') - then show ?thesis - proof (cases "fmlookup (storage st'') (address env)") - case None - with * ** *** KStoptr Storage STArray LStackloc Memory Some KMemptr show ?thesis by (simp split:if_split_asm) - next - case s2: (Some s) - then show ?thesis - proof (cases "cps2m p p' x t s (memory st'')") - case None - with * ** *** KStoptr Storage STArray LStackloc Memory Some KMemptr s2 show ?thesis by (simp split:if_split_asm) - next - case s3: (Some m) - with * ** *** KStoptr Storage STArray LStackloc Memory Some KMemptr s2 have "st6' = st'' \memory := m\" by (auto split:if_split_asm) - with f3 show ?thesis using asm by (simp add:frame_def) - qed - qed - next - case m2: (KStoptr x3) - with * ** *** KStoptr Storage STArray LStackloc Memory Some show ?thesis by (simp split:if_split_asm) - qed - qed - next - case st2: (Storage x4) - with f3 * ** *** KStoptr Storage STArray LStackloc show ?thesis by (auto simp add:frame_def split:if_split_asm) - qed - next - case (LStoreloc l') - then show ?thesis - proof (cases "fmlookup (storage st'') (address env)") - case None - with * ** *** KStoptr Storage STArray LStoreloc show ?thesis by (auto split:if_split_asm) - next - case (Some s) - then show ?thesis - proof (cases "copy p l' x t s") - case None - with * ** *** KStoptr Storage STArray LStoreloc Some show ?thesis by (auto split:if_split_asm) - next - case s2: (Some s') - with * ** *** KStoptr Storage STArray LStoreloc Some s2 have "st6' = st'' \storage := fmupd (address env) s' (storage st'')\" by (auto split:if_split_asm) - with f3 show ?thesis using asm by (simp add:frame_def) - qed - qed - next - case (LMemloc l') - then show ?thesis - proof (cases "fmlookup (storage st'') (address env)") - case None - with * ** *** KStoptr Storage STArray LMemloc show ?thesis by (auto split:if_split_asm) - next - case (Some s) - then show ?thesis - proof (cases "cps2m p l' x t s (memory st'')") - case None - with * ** *** KStoptr Storage STArray LMemloc Some show ?thesis by (auto split:if_split_asm) - next - case s2: (Some m) - with * ** *** KStoptr Storage STArray LMemloc Some s2 have "st6' = st'' \memory := m\" by (auto split:if_split_asm) - with f3 show ?thesis using asm by (simp add:frame_def) - qed - qed - qed - next - case (STMap t t') - with * ** KStoptr Storage obtain l' rt st'' where ***: "lexp lv e\<^sub>p env cd st' = Normal ((LStackloc l',rt), st'')" by (auto split:if_split_asm result.split_asm LType.split_asm) - with KStoptr Storage STMap have "\rv5 st5' (ev5'::Environment) bal. - frame bal st' \ - local.lexp lv e\<^sub>p env cd st' = Normal (rv5, st5') \ - frame bal st5'" using asm 0 37(6)[OF a1 **] by auto - with f2 *** have f3: "frame bal st''" by blast - moreover from * ** *** KStoptr Storage STMap have "st6' = st'' \stack := updateStore l' (KStoptr p) (stack st'')\" by (auto split:if_split_asm) - ultimately show ?thesis using asm f3 by (simp add:frame_def) - next - case (STValue x2) - with * ** KStoptr Storage show ?thesis by (simp split:if_split_asm) - qed - next - case (Memory x4) - with * ** KStoptr show ?thesis by (simp split:if_split_asm) - qed - qed - qed - qed - qed - next - show "?RHS" (is "?LHS \ ?RHS") - proof - assume "address env = STR ''Victim''" - show ?RHS (is "?A \ (?B \ ?C)") - proof (rule conj3) - show ?A (is "\s val bal x. ?LHS s val bal x") - proof (rule allI[OF allI[OF allI[OF allI]]]) - fix s val bal x - show "?LHS s val bal x" (is "?LHS \ ?RHS") - proof - assume ?LHS - then show ?RHS by simp - qed - qed - next - show ?B (is "\s bal x. ?LHS s bal x") - proof (rule allI[OF allI[OF allI]]) - fix s bal x - show "?LHS s bal x" (is "?LHS \ ?RHS") - proof - assume ?LHS - then show ?RHS by simp - qed - qed - next - show ?C (is "\s bal. ?LHS s bal") - proof (rule allI[OF allI]) - fix s bal - show "?LHS s bal" (is "?LHS \ ?RHS") - proof - assume ?LHS - then show ?RHS by simp - qed - qed - qed - qed - qed - qed - qed - qed -next - case (38 s1 s2 e\<^sub>p e cd st) - show ?case (is "?LHS \ ?RHS") - proof - assume 0: "fmlookup e\<^sub>p STR ''Victim'' = Some (victim, SKIP)" - show ?RHS (is "\st6'. ?RHS st6'") - proof - fix st6' - show "?RHS st6'" (is "?LHS \ ?RHS") - proof - assume *: "stmt (COMP s1 s2) e\<^sub>p e cd st = Normal ((), st6')" - show ?RHS (is "?LHS \ ?RHS") - proof - show "?LHS" - proof - assume asm: "address e \ STR ''Victim''" - show "\bal. frame bal st \ frame bal st6'" - proof - fix bal - show "frame bal st \ frame bal st6'" - proof - assume "frame bal st" - with * have a1: "(applyf (costs (COMP s1 s2) e\<^sub>p e cd) \ (\g. assert Gas (\st. gas st \ g) (modify (\st. st\gas := gas st - g\)))) st = - Normal ((), st\gas:=gas st - costs (COMP s1 s2) e\<^sub>p e cd st\)" - and f1: "frame bal (st\gas:=gas st - costs (COMP s1 s2) e\<^sub>p e cd st\)" by (auto simp add:frame_def) - then have "\rv4 st4' bal. - frame bal (st\gas := gas st - costs (COMP s1 s2) e\<^sub>p e cd st\) \ - stmt s1 e\<^sub>p e cd (st\gas := gas st - costs (COMP s1 s2) e\<^sub>p e cd st\) = Normal (rv4, st4') \ - frame bal st4'" using asm 0 38(1) by (simp add:frame_def) - moreover from * obtain st' where **: "stmt s1 e\<^sub>p e cd (st\gas:=gas st - costs (COMP s1 s2) e\<^sub>p e cd st\) = Normal ((), st')" by (auto split:if_split_asm result.split_asm) - ultimately have f2: "frame bal st'" using f1 by blast - - have "\rv4 st4' bal. - frame bal st' \ - stmt s2 e\<^sub>p e cd st' = Normal (rv4, st4') \ - frame bal st4'" using asm 0 38(2)[OF a1 **] by (simp add:frame_def) - moreover from * ** obtain st'' where ***: "stmt s2 e\<^sub>p e cd st' = Normal ((), st'')" by (auto split:if_split_asm result.split_asm) - ultimately have f3: "frame bal st''" using f2 by blast - - from a1 * ** *** have "st6' = st''" by (simp split:if_split_asm) - with f3 asm show "frame bal st6'" by simp - qed - qed - qed - next - show "?RHS" (is "?LHS \ ?RHS") - proof - assume ad: "address e = STR ''Victim''" - show ?RHS (is "?A \ ?B \ ?C") - proof (rule conj3) - show ?A (is "\s val bal x. ?LHS s val bal x") - proof (rule allI[OF allI[OF allI[OF allI]]]) - fix s val bal x - show "?LHS s val bal x" (is "?LHS \ ?RHS") - proof - assume ?LHS - then show ?RHS by simp - qed - qed - next - show ?B (is "\s bal x. ?LHS s bal x") - proof (rule allI[OF allI[OF allI]]) - fix s bal x - show "?LHS s bal x" (is "?LHS \ ?RHS") - proof - assume ?LHS(is "?A1 \ ?A2 \ ?A3 \ ?A4 \ ?A5 \ ?A6") - then have ?A1 and ?A2 and ?A3 and ?A4 and ?A5 and ?A6 by auto - with * have c1: "gas st > costs comp e\<^sub>p e cd st" by (auto split:if_split_asm) - with `?A1` * obtain st'' where 00: "stmt assign e\<^sub>p e cd (st\gas := gas st - costs comp e\<^sub>p e cd st\) = Normal((), st'')" by (auto split:result.split_asm) - moreover from `?A2` have "fmlookup (storage (st\gas := gas st - costs comp e\<^sub>p e cd st\)) (STR ''Victim'') = Some s" by simp - moreover from `?A2` have "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts (st\gas := gas st - costs comp e\<^sub>p e cd st\)) (STR ''Victim'')) - (SUMM s) \ bal \ bal \ 0" by simp - moreover from `?A3` have "POS s" by simp - moreover from `?A6` have "accessStore x (stack (st\gas := gas st - costs comp e\<^sub>p e cd st\)) = Some (KValue (accessStorage (TUInt 256) (sender e + (STR ''.'' + STR ''balance'')) s))" by simp - ultimately obtain s'' where "fmlookup (storage st'') (STR ''Victim'') = Some s''" - and "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st'') (STR ''Victim'')) - (SUMM s'' + ReadL\<^sub>i\<^sub>n\<^sub>t (accessStorage (TUInt 256) (sender e + (STR ''.'' + STR ''balance'')) s)) \ bal \ bal \ 0" - and **: "accessStore x (stack st'') = Some (KValue (accessStorage (TUInt 256) (sender e + (STR ''.'' + STR ''balance'')) s))" - and "POS s''" - using secureassign[OF 00 _ ad `?A5`] that by blast - moreover from c1 `?A1` * 00 obtain st''' where ***: "stmt transfer e\<^sub>p e cd st'' = Normal((), st''')" and "st6' = st'''" by auto - - moreover from `?A1` 00 have x1: "stmt s1 e\<^sub>p e cd (st\gas := gas st - costs (COMP s1 s2) e\<^sub>p e cd st\) = Normal((), st'')" by simp - moreover from * have x2: "(applyf (costs (COMP s1 s2) e\<^sub>p e cd) \ (\g. assert Gas - (\st. gas st \ g) - (modify (\st. st\gas := gas st - g\)))) - st = Normal ((), st\gas:=gas st - costs (COMP s1 s2) e\<^sub>p e cd st\)" by (simp split: if_split_asm) - ultimately show "\s'. fmlookup (storage st6') (STR ''Victim'') = Some s' - \ ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st6') (STR ''Victim'')) - (SUMM s') \ bal \ bal \ 0 \ POS s'" - using 38(2)[OF x2 x1] `?A1` `?A4` ad 0 ** `?A6` by simp - qed - qed - next - show ?C (is "\s bal. ?LHS s bal") - proof (rule allI[OF allI]) - fix s bal - show "?LHS s bal" (is "?LHS \ ?RHS") - proof - assume ?LHS - then show ?RHS by simp - qed - qed - qed - qed - qed - qed - qed - qed -next - case (39 ex s1 s2 e\<^sub>p e cd st) - show ?case (is "?LHS \ ?RHS") - proof - assume 0: "fmlookup e\<^sub>p STR ''Victim'' = Some (victim, SKIP)" - show ?RHS (is "\st6'. ?RHS st6'") - proof - fix st6' - show "?RHS st6'" (is "?LHS \ ?RHS") - proof - assume *: "stmt (ITE ex s1 s2) e\<^sub>p e cd st = Normal ((), st6')" - show ?RHS (is "?LHS \ ?RHS") - proof - show "?LHS" - proof - assume asm: "address e \ STR ''Victim''" - show "\bal. frame bal st \ frame bal st6'" - proof - fix bal - show "frame bal st \ frame bal st6'" - proof - assume "frame bal st" - with * have a1: "(applyf (costs (ITE ex s1 s2) e\<^sub>p e cd) \ (\g. assert Gas (\st. gas st \ g) (modify (\st. st\gas := gas st - g\)))) st = - Normal ((), st\gas:=gas st - costs (ITE ex s1 s2) e\<^sub>p e cd st\)" - and f1: "frame bal (st\gas:=gas st - costs (ITE ex s1 s2) e\<^sub>p e cd st\)" by (auto simp add:frame_def) - - from * obtain b st' where **: "expr ex e\<^sub>p e cd (st\gas:=gas st - costs (ITE ex s1 s2) e\<^sub>p e cd st\) = Normal ((KValue b, Value TBool), st')" by (auto split:if_split_asm result.split_asm prod.split_asm Stackvalue.split_asm Type.split_asm Types.split_asm) - moreover from asm have "\rv4 st4' bal. - frame bal (st\gas := gas st - costs (ITE ex s1 s2) e\<^sub>p e cd st\) \ - expr ex e\<^sub>p e cd (st\gas := gas st - costs (ITE ex s1 s2) e\<^sub>p e cd st\) = Normal (rv4, st4') \ - frame bal st4'" using 0 39(1)[OF a1] by (simp add:frame_def) - ultimately have f2: "frame bal st'" using f1 by blast - - show "frame bal st6'" - proof (cases "b = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True") - case True - then have "\st6' bal. - frame bal st' \ - local.stmt s1 e\<^sub>p e cd st' = Normal ((), st6') \ - frame bal st6'" using asm 0 39(2)[OF a1 **, of "KValue b" "Value TBool" b TBool] by (simp add:frame_def) - moreover from * ** True obtain st'' where ***: "stmt s1 e\<^sub>p e cd st' = Normal ((), st'')" by (auto split:if_split_asm result.split_asm) - ultimately have "frame bal st''" using f2 by blast - moreover from a1 * ** True *** have "st6' = st''" by (simp split:if_split_asm) - ultimately show "frame bal st6'" using asm by simp - next - case False - then have "\st6' bal. - frame bal st' \ - local.stmt s2 e\<^sub>p e cd st' = Normal ((), st6') \ - frame bal st6'" using 0 asm 39(3)[OF a1 **, of "KValue b" "Value TBool" b TBool] by (simp add:frame_def) - moreover from * ** False obtain st'' where ***: "stmt s2 e\<^sub>p e cd st' = Normal ((), st'')" by (auto split:if_split_asm result.split_asm) - ultimately have "frame bal st''" using f2 by blast - moreover from a1 * ** False *** have "st6' = st''" by (simp split:if_split_asm) - ultimately show "frame bal st6'" using asm by simp - qed - qed - qed - qed - next - show "?RHS" (is "?LHS \ ?RHS") - proof - assume "address e = STR ''Victim''" - show ?RHS (is "?A \ (?B \ ?C)") - proof (rule conj3) - show ?A (is "\s val bal x. ?LHS s val bal x") - proof (rule allI[OF allI[OF allI[OF allI]]]) - fix s val bal x - show "?LHS s val bal x" (is "?LHS \ ?RHS") - proof - assume ?LHS - then show ?RHS by simp - qed - qed - next - show ?B (is "\s bal x. ?LHS s bal x") - proof (rule allI[OF allI[OF allI]]) - fix s bal x - show "?LHS s bal x" (is "?LHS \ ?RHS") - proof - assume ?LHS - then show ?RHS by simp - qed - qed - next - show ?C (is "\s bal. ?LHS s bal") - proof (rule allI[OF allI]) - fix s bal - show "?LHS s bal" (is "?LHS \ ?RHS") - proof - assume ?LHS - then show ?RHS by simp - qed - qed - qed - qed - qed - qed - qed - qed -next - case (40 ex s0 e\<^sub>p e cd st) - show ?case (is "?LHS \ ?RHS") - proof - assume 0: "fmlookup e\<^sub>p STR ''Victim'' = Some (victim, SKIP)" - show ?RHS (is "\st6'. ?RHS st6'") - proof - fix st6' - show "?RHS st6'" (is "?LHS \ ?RHS") - proof - assume *: "stmt (WHILE ex s0) e\<^sub>p e cd st = Normal ((), st6')" - show ?RHS (is "?LHS \ ?RHS") - proof - show "?LHS" - proof - assume asm: "address e \ STR ''Victim''" - show "\bal. frame bal st \ frame bal st6'" - proof - fix bal - show "frame bal st \ frame bal st6'" - proof - assume "frame bal st" - with * have a1: "(applyf (costs (WHILE ex s0) e\<^sub>p e cd) \ (\g. assert Gas (\st. gas st \ g) (modify (\st. st\gas := gas st - g\)))) st = - Normal ((), st\gas:=gas st - costs (WHILE ex s0) e\<^sub>p e cd st\)" - and f1: "frame bal (st\gas:=gas st - costs (WHILE ex s0) e\<^sub>p e cd st\)" by (auto simp add:frame_def) - - from * obtain b st' where **: "expr ex e\<^sub>p e cd (st\gas:=gas st - costs (WHILE ex s0) e\<^sub>p e cd st\) = Normal ((KValue b, Value TBool), st')" by (auto split:if_split_asm result.split_asm prod.split_asm Stackvalue.split_asm Type.split_asm Types.split_asm) - moreover from asm have "\rv4 st4' bal. - frame bal (st\gas := gas st - costs (WHILE ex s0) e\<^sub>p e cd st\) \ - expr ex e\<^sub>p e cd (st\gas := gas st - costs (WHILE ex s0) e\<^sub>p e cd st\) = Normal (rv4, st4') \ - frame bal st4'" using 0 40(1)[OF a1] by (simp add:frame_def) - ultimately have f2: "frame bal st'" using f1 by blast - - show "frame bal st6'" - proof (cases "b = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True") - case True - then have "\st6' bal. - frame bal st' \ - local.stmt s0 e\<^sub>p e cd st' = Normal ((), st6') \ - frame bal st6'" using 0 asm 40(2)[OF a1 **, of "KValue b" "Value TBool" b TBool] by (simp add:frame_def) - moreover from * ** True obtain st'' where ***: "stmt s0 e\<^sub>p e cd st' = Normal ((), st'')" by (auto split:if_split_asm result.split_asm) - ultimately have f3: "frame bal st''" using f2 by blast - - have "\st6' bal. - frame bal st'' \ - local.stmt (WHILE ex s0) e\<^sub>p e cd st'' = Normal ((), st6') \ - frame bal st6'" using 0 asm 40(3)[OF a1 ** _ _ _ _ True ***] by (simp add:frame_def) - moreover from * ** True *** obtain st''' where ****: "stmt (WHILE ex s0) e\<^sub>p e cd st'' = Normal ((), st''')" by (auto split:if_split_asm result.split_asm) - ultimately have "frame bal st'''" using f3 by blast - - moreover from a1 * ** True *** **** have "st6' = st'''" by (simp split:if_split_asm) - ultimately show "frame bal st6'" using asm by simp - next - case False - then show "frame bal st6'" using * ** f1 f2 asm by (simp split:if_split_asm) - qed - qed - qed - qed - next - show "?RHS" (is "?LHS \ ?RHS") - proof - assume "address e = STR ''Victim''" - show ?RHS (is "?A \ (?B \ ?C)") - proof (rule conj3) - show ?A (is "\s val bal x. ?LHS s val bal x") - proof (rule allI[OF allI[OF allI[OF allI]]]) - fix s val bal x - show "?LHS s val bal x" (is "?LHS \ ?RHS") - proof - assume ?LHS - then show ?RHS by simp - qed - qed - next - show ?B (is "\s bal x. ?LHS s bal x") - proof (rule allI[OF allI[OF allI]]) - fix s bal x - show "?LHS s bal x" (is "?LHS \ ?RHS") - proof - assume ?LHS - then show ?RHS by simp - qed - qed - next - show ?C (is "\s bal. ?LHS s bal") - proof (rule allI[OF allI]) - fix s bal - show "?LHS s bal" (is "?LHS \ ?RHS") - proof - assume ?LHS - then show ?RHS by simp - qed - qed - qed - qed - qed - qed - qed - qed -next - case (41 i xe e\<^sub>p e cd st) - show ?case (is "?LHS \ ?RHS") - proof - assume 0: "fmlookup e\<^sub>p STR ''Victim'' = Some (victim, SKIP)" - show ?RHS (is "\st6'. ?RHS st6'") - proof - fix st' - show "?RHS st'" (is "?LHS \ ?RHS") - proof - assume *: "stmt (INVOKE i xe) e\<^sub>p e cd st = Normal ((), st')" - show ?RHS (is "?LHS \ ?RHS") - proof - show "?LHS" - proof - assume asm: "address e \ STR ''Victim''" - show "\bal. frame bal st \ frame bal st'" - proof - fix bal - show "frame bal st \ frame bal st'" - proof - assume ff: "frame bal st" - moreover from * have a1: "(applyf (costs (INVOKE i xe) e\<^sub>p e cd) \ (\g. assert Gas (\st. gas st \ g) (modify (\st. st\gas := gas st - g\)))) st = Normal ((), st\gas := gas st - costs (INVOKE i xe) e\<^sub>p e cd st\)" by auto - moreover from * obtain ct bla where **: "fmlookup e\<^sub>p (address e) = Some (ct, bla)" - by (auto split:if_split_asm option.split_asm) - moreover from * ** obtain fp f where ***: "fmlookup ct i = Some (Method (fp, f, None))" - by (auto split:if_split_asm option.split_asm Member.split_asm) - moreover define e' where "e' = ffold_init ct (emptyEnv (address e) (sender e) (svalue e)) (fmdom ct)" - moreover from * ** *** obtain e'' cd' st'' st''' where ****: "load False fp xe e\<^sub>p e' emptyStore (st\gas:=gas st - (costs (INVOKE i xe) e\<^sub>p e cd st), stack:=emptyStore\) e cd (st\gas:=gas st - (costs (INVOKE i xe) e\<^sub>p e cd st)\) = Normal ((e'', cd', st''), st''')" - using e'_def by (auto split:if_split_asm result.split_asm) - moreover from * **** have f1: "frame bal st''" and ad: "address e' = address e''" - using asm ff 0 41(1)[OF a1 ** _ *** _ _ _ _ e'_def, of bla "(fp, f, None)" fp "(f, None)" f None] by (auto simp add:frame_def) - moreover from e'_def have ad2: "address e = address e'" using ffold_init_ad_same[of ct "(emptyEnv (address e) (sender e) (svalue e))" "(fmdom ct)" e'] by simp - moreover from * ** *** **** e'_def obtain st'''' where *****: "stmt f e\<^sub>p e'' cd' st'' = Normal ((), st'''')" by (auto split:if_split_asm result.split_asm) - ultimately have "st' = st''''\stack:=stack st''', memory := memory st'''\" using * apply safe by simp - moreover from f1 ad ad2 asm ***** have f2:"frame bal st''''" - using 41(2)[OF a1 ** _ *** _ _ _ _ e'_def _ ****] using 0 * by (auto simp add:frame_def) - ultimately show "frame bal st'" by (simp add:frame_def) - qed - qed - qed - next - show "?RHS" (is "?LHS \ ?RHS") - proof - assume "address e = STR ''Victim''" - show ?RHS (is "?A \ (?B \ ?C)") - proof (rule conj3) - show ?A (is "\s val bal x. ?LHS s val bal x") - proof (rule allI[OF allI[OF allI[OF allI]]]) - fix s val bal x - show "?LHS s val bal x" (is "?LHS \ ?RHS") - proof - assume ?LHS - then show ?RHS by simp - qed - qed - next - show ?B (is "\s bal x. ?LHS s bal x") - proof (rule allI[OF allI[OF allI]]) - fix s bal x - show "?LHS s bal x" (is "?LHS \ ?RHS") - proof - assume ?LHS - then show ?RHS by simp - qed - qed - next - show ?C (is "\s bal. ?LHS s bal") - proof (rule allI[OF allI]) - fix s bal - show "?LHS s bal" (is "?LHS \ ?RHS") - proof - assume ?LHS - then show ?RHS by simp - qed - qed - qed - qed - qed - qed - qed - qed -next - case (42 ad i xe val e\<^sub>p e cd st) - show ?case (is "?LHS \ ?RHS") - proof - assume 0: "fmlookup e\<^sub>p STR ''Victim'' = Some (victim, SKIP)" - show ?RHS (is "\st6'. ?RHS st6'") - proof - fix st' - show "?RHS st'" (is "?LHS \ ?RHS") - proof - assume *: "stmt (EXTERNAL ad i xe val) e\<^sub>p e cd st = Normal ((), st')" - show ?RHS (is "?LHS \ ?RHS") - proof - show "?LHS" - proof - assume asm: "address e \ STR ''Victim''" - show "\bal. frame bal st \ frame bal st'" - proof - fix bal - show "frame bal st \ frame bal st'" - proof - assume ff: "frame bal st" - moreover from * have a1: "(applyf (costs (EXTERNAL ad i xe val) e\<^sub>p e cd) \ (\g. assert Gas (\st. gas st \ g) (modify (\st. st\gas := gas st - g\)))) st = Normal ((), st\gas := gas st - costs (EXTERNAL ad i xe val) e\<^sub>p e cd st\)" by auto - moreover from * obtain adv st'' where **: "expr ad e\<^sub>p e cd (st\gas:=gas st - (costs (EXTERNAL ad i xe val) e\<^sub>p e cd st)\) = Normal ((KValue adv, Value TAddr), st'')" - by (auto split:if_split_asm result.split_asm Stackvalue.split_asm Type.split_asm Types.split_asm) - moreover from * ** ff have f1: "frame bal st''" using asm 0 42(1) by (simp add:frame_def split:if_split_asm) - moreover from * ** obtain ct fb where ***: "fmlookup e\<^sub>p adv = Some (ct, fb)" - by (auto split:if_split_asm option.split_asm) - moreover from * ** *** f1 obtain v t st''' where ****: "expr val e\<^sub>p e cd st'' = Normal ((KValue v, Value t), st''')" - by (auto split:if_split_asm result.split_asm Stackvalue.split_asm Type.split_asm) - moreover from **** f1 have "frame bal st'''" using asm 42(2)[OF a1 ** _ _ _ _ ***] 0 by (simp split:if_split_asm) - then have f2: "frame bal (st'''\stack := emptyStore, memory := emptyStore\)" by (simp add:frame_def) - show "frame bal st'" - proof (cases "fmlookup ct i") - case None - with * ** *** **** obtain acc where trans: "Accounts.transfer (address e) adv v (accounts st''') = Some acc" by (auto split:if_split_asm option.split_asm) - with * ** *** **** None obtain st'''' where *****: "stmt fb e\<^sub>p (emptyEnv adv (address e) v) cd (st'''\accounts := acc,stack:=emptyStore, memory:=emptyStore\) = Normal ((), st'''')" - by (auto split:if_split_asm result.split_asm) - moreover have f4: "frame bal (st''''\stack:=stack st'', memory := memory st''\)" - proof (cases "adv = STR ''Victim''") - case True - with 0 *** have "fb = SKIP" by simp - moreover from f2 have "frame bal (st'''\accounts := acc,stack:=emptyStore, memory:=emptyStore\)" using transfer_frame[OF trans] asm by (simp add:frame_def) - ultimately show ?thesis using ***** by (auto split:if_split_asm simp add:frame_def) - next - case False - moreover from f2 have "frame bal (st'''\accounts := acc,stack:=emptyStore, memory:=emptyStore\)" using transfer_frame[OF trans] asm by (simp add:frame_def) - then have "frame bal st''''" using f2 0 42(5)[OF a1 ** _ _ _ _ *** _ **** _ _ _ None _ trans, of "KValue adv" "Value TAddr" TAddr fb "KValue v" "Value t" t st''' st''' st'''] asm ***** False by (auto simp add:frame_def) - then show ?thesis by (simp add:frame_def) - qed - ultimately show "frame bal st'" using a1 * ** *** **** None trans by (auto simp add:frame_def) - next - case (Some a) - with * ** *** **** obtain fp f where *****: "fmlookup ct i = Some (Method (fp, f, None))" - by (auto split:if_split_asm option.split_asm Member.split_asm) - moreover define e' where e'_def: "e' = ffold_init ct (emptyEnv adv (address e) v) (fmdom ct)" - moreover from * ** *** ***** **** obtain e'' cd' st'''' st''''' where *******: "load True fp xe e\<^sub>p e' emptyStore (st'''\stack:=emptyStore, memory:=emptyStore\) e cd st''' = Normal ((e'', cd', st''''), st''''')" - using e'_def by (auto split:if_split_asm result.split_asm option.split_asm) - moreover from e'_def have ad2: "address e' = adv" and send2: "sender e' = address e" and sval2: "svalue e' = v" using ffold_init_ad_same[of ct "(emptyEnv adv (address e) v)" "(fmdom ct)" e'] by auto - moreover from * ** *** ***** **** ******* e'_def obtain acc where trans: "Accounts.transfer (address e) adv v (accounts st'''') = Some acc" by (auto split:if_split_asm option.split_asm) - then have ******: "Accounts.transfer (address e) adv v (accounts st'''') = Some acc" by (auto split:if_split_asm option.split_asm) - moreover from * ** *** ***** **** ****** ******* obtain st'''''' where ********: "stmt f e\<^sub>p e'' cd' (st''''\accounts := acc\) = Normal ((), st'''''')" - using e'_def by (auto split:if_split_asm result.split_asm) - moreover have f4: "frame bal st''''''" - proof (cases "adv = STR ''Victim''") - case True - with 0 *** have ct_def: "ct = victim" by simp - - moreover have - "(\(ev::Environment) cda st st' bal. - local.load True fp xe e\<^sub>p e' emptyStore (st'''\stack := emptyStore, memory := emptyStore\) e cd st''' = Normal ((ev, cda, st), st') \ - (frame bal (st'''\stack := emptyStore, memory := emptyStore\) \ frame bal st) \ - (frame bal st''' \ frame bal st') \ address e' = address ev \ sender e' = sender ev \ svalue e' = svalue ev)" - using 0 42(3)[OF a1 ** _ _ _ _ *** _ **** _ _ _ ***** _ _ _ _ e'_def] asm by simp - with f2 ******* have f3: "frame bal st''''" and ad1: "address e' = address e''" and send1: "sender e' = sender e''" and sval1: "svalue e' = svalue e''" by auto - - from ct_def ***** consider (withdraw) "i = STR ''withdraw''" | (deposit) "i = STR ''deposit''" using victim_def fmap_of_list_SomeD[of "[(STR ''balance'', Var (STMap TAddr (STValue (TUInt 256)))), (STR ''deposit'', Method ([], deposit, None)), (STR ''withdraw'', Method ([], keep, None))]"] by auto - then show ?thesis - proof cases - case withdraw - moreover have "fmlookup victim (STR ''withdraw'') = Some (Method ([], keep, None))" using victim_def by eval - ultimately have "f=keep" and "fp = []" using *** ***** True 0 by auto - with ******** have *********: "stmt keep e\<^sub>p e'' cd' (st''''\accounts := acc\) = Normal ((), st'''''')" by simp - have "fmlookup (denvalue e'') STR ''balance'' = Some (Storage (STMap TAddr (STValue (TUInt 256))), Storeloc STR ''balance'')" - proof - - from victim_def have some: "fmlookup victim (STR ''balance'') = Some (Var (STMap TAddr (STValue (TUInt 256))))" by eval - with ct_def have "fmlookup ct (STR ''balance'') = Some (Var (STMap TAddr (STValue (TUInt 256))))" by simp - moreover have "STR ''balance'' |\| fmdom (denvalue (emptyEnv adv (address e) v))" by simp - moreover from ct_def some have "STR ''balance'' |\| fmdom ct" using fmdomI by simp - ultimately have "fmlookup (denvalue e') STR ''balance'' = Some (Storage (STMap TAddr (STValue (TUInt 256))), Storeloc STR ''balance'')" using e'_def ffold_init_fmap[of ct "STR ''balance''" "(STMap TAddr (STValue (TUInt 256)))" "(emptyEnv adv (address e) v)" "(fmdom ct)"] by simp - moreover have "e'' = e'" - proof (cases "xe=[]") - case True - with ******* `fp = []` show ?thesis by simp - next - case False - then obtain xx xe' where "xe = xx # xe'" using list.exhaust by auto - with ******* `fp = []` show ?thesis by simp - qed - ultimately show ?thesis by simp - qed - - moreover from ad1 ad2 True have ad: "address e'' = STR ''Victim''" by simp - moreover from ad send1 send2 asm have "sender e'' \ address e''" by simp - moreover from f3 have f4: "frame bal (st''''\accounts := acc\)" using transfer_frame[OF ******] asm by simp - then obtain s'''' where "fmlookup (storage (st''''\accounts := acc\)) (STR ''Victim'') = Some s'''' \ ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts (st''''\accounts := acc\)) (STR ''Victim'')) - (SUMM s'''') \ bal \ bal \ 0 \ POS s''''" by (auto simp add:frame_def) - ultimately have "(\s'. fmlookup (storage st'''''') (STR ''Victim'') = Some s' - \ ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st'''''') (STR ''Victim'')) - (SUMM s') \ bal \ bal \ 0 \ POS s')" - using 0 ******** `f=keep` 42(4)[OF a1 ** _ _ _ _ *** _ **** _ _ _ ***** _ _ _ _ e'_def _ ******* _ _ _ trans, of "KValue adv" "Value TAddr" TAddr fb "KValue v" "Value t" t "(fp, f, None)" "(f, None)" f None e'' "(cd', st'''')" cd' st''''' st''''' "()" "st''''\accounts := acc\"] apply safe by auto - then show ?thesis by (simp add:frame_def) - next - case deposit - moreover from f2 have "frame bal (st'''\stack:=emptyStore, memory:=emptyStore\)" using transfer_frame[OF ******] asm by simp - moreover have "fmlookup victim (STR ''deposit'') = Some (Method ([], deposit, None))" using victim_def by eval - ultimately have "f=deposit" and "fp = []" using *** ***** True 0 by auto - with ******** have *: "stmt deposit e\<^sub>p e'' cd' (st''''\accounts := acc\) = Normal ((), st'''''')" by simp - - have f4: "frame (bal + ReadL\<^sub>i\<^sub>n\<^sub>t v) (st''''\accounts := acc\)" and ad1: "address e' = address e''" and send1: "sender e' = sender e''" and sval1: "svalue e' = svalue e''" - proof - - have "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance acc (STR ''Victim'')) = ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st'''') (STR ''Victim'')) + ReadL\<^sub>i\<^sub>n\<^sub>t v" using transfer_add[OF ******] asm True by simp - moreover have "ReadL\<^sub>i\<^sub>n\<^sub>t v \ 0" using transfer_val1[OF ******] by simp - ultimately have "frame (bal + ReadL\<^sub>i\<^sub>n\<^sub>t v) (st''''\accounts := acc\)" using f3 by (auto simp add:frame_def) - then show "frame (bal + ReadL\<^sub>i\<^sub>n\<^sub>t v) (st''''\accounts := acc\)" and "address e' = address e''" and "sender e' = sender e''" and "svalue e' = svalue e''" using f2 0 42(3)[OF a1 ** _ _ _ _ *** _ **** _ _ _ ***** _ _ _ _ e'_def, of "KValue adv" "Value TAddr" TAddr fb "KValue v" "Value t"] asm ******* by (auto simp add:frame_def) - qed - moreover from sval1 sval2 have "v = svalue e''" by simp - ultimately have "frame (bal + ReadL\<^sub>i\<^sub>n\<^sub>t (svalue e'')) (st''''\accounts := acc\)" by simp - then obtain s''''' where II: "INV (st''''\accounts := acc\) s''''' (SUMM s''''') (bal + ReadL\<^sub>i\<^sub>n\<^sub>t (svalue e''))" and III:"POS s'''''" by (auto simp add:frame_def) - then have s'''''_def: "fmlookup (storage (st''''\accounts := acc\)) STR ''Victim'' = Some s'''''" by simp - - have yyy: "fmlookup (denvalue e'') STR ''balance'' = Some (Storage (STMap TAddr (STValue (TUInt 256))), Storeloc STR ''balance'')" - proof - - from victim_def have some: "fmlookup victim (STR ''balance'') = Some (Var (STMap TAddr (STValue (TUInt 256))))" by eval - with ct_def have "fmlookup ct (STR ''balance'') = Some (Var (STMap TAddr (STValue (TUInt 256))))" by simp - moreover have "STR ''balance'' |\| fmdom (denvalue (emptyEnv adv (address e) v))" by simp - moreover from ct_def some have "STR ''balance'' |\| fmdom ct" using fmdomI by simp - ultimately have "fmlookup (denvalue e') STR ''balance'' = Some (Storage (STMap TAddr (STValue (TUInt 256))), Storeloc STR ''balance'')" using e'_def ffold_init_fmap[of ct "STR ''balance''" "(STMap TAddr (STValue (TUInt 256)))" "(emptyEnv adv (address e) v)" "(fmdom ct)"] by simp - moreover have "e'' = e'" - proof (cases "xe=[]") - case True - with ******* `fp = []` show ?thesis by simp - next - case False - then obtain xx xe' where "xe = xx # xe'" using list.exhaust by auto - with ******* `fp = []` show ?thesis by simp - qed - ultimately show ?thesis by simp - qed - - from asm True have "address e \ (STR ''Victim'')" by simp - then have "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st'''') (STR ''Victim'')) + ReadL\<^sub>i\<^sub>n\<^sub>t v < 2^256" using transfer_val2[OF ******] True by simp - moreover from `address e \ (STR ''Victim'')` have "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance acc (STR ''Victim'')) = ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st'''') (STR ''Victim'')) + ReadL\<^sub>i\<^sub>n\<^sub>t v" using transfer_add[OF ******] (******) True by simp - ultimately have abc: "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts (st''''\accounts := acc\)) (STR ''Victim'')) < 2^256" by simp - - from II have "fmlookup (storage (st''''\accounts := acc\)) (STR ''Victim'') = Some s''''' \ ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts (st''''\accounts := acc\)) (STR ''Victim'')) - (SUMM s''''') \ bal + ReadL\<^sub>i\<^sub>n\<^sub>t (svalue e'') \ bal + ReadL\<^sub>i\<^sub>n\<^sub>t (svalue e'') \ 0" by (auto) - moreover have "ReadL\<^sub>i\<^sub>n\<^sub>t (accessStorage (TUInt 256) (sender e'' + (STR ''.'' + STR ''balance'')) s''''') + \svalue e''\ < 2 ^ 256 \ - ReadL\<^sub>i\<^sub>n\<^sub>t (accessStorage (TUInt 256) (sender e'' + (STR ''.'' + STR ''balance'')) s''''') + \svalue e''\ \ 0" - proof (cases "fmlookup s''''' (sender e'' + (STR ''.'' + STR ''balance'')) = None") - case True - hence "(accessStorage (TUInt 256) (sender e'' + (STR ''.'' + STR ''balance'')) s''''') = ShowL\<^sub>i\<^sub>n\<^sub>t 0" by simp - moreover have "(\svalue e''\::int) < 2 ^ 256" - proof - - from II have "bal + \svalue e''\ + SUMM s''''' \ ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts (st''''\accounts := acc\)) (STR ''Victim''))" by simp - moreover have "0 \ SUMM s'''''" - using III sum_nonneg[of "{(ad,x). fmlookup s''''' (ad + (STR ''.'' + STR ''balance'')) = Some x}" "\(ad,x). ReadL\<^sub>i\<^sub>n\<^sub>t x"] by auto - ultimately have "bal + \svalue e''\ \ ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts (st''''\accounts := acc\)) (STR ''Victim''))" by simp - moreover from ff have "bal\0" by (auto simp add:frame_def) - ultimately show "(\svalue e''\::int) < 2 ^ 256" using abc by simp - qed - moreover have "ReadL\<^sub>i\<^sub>n\<^sub>t v \ 0" using transfer_val1[OF ******] by simp - with `svalue e' = v` sval1 have "(\svalue e''\::int) \ 0" by simp - ultimately show ?thesis using Read_ShowL_id by simp - next - case False - then obtain x where x_def: "fmlookup s''''' (sender e'' + (STR ''.'' + STR ''balance'')) = Some x" by auto - moreover from II have "bal + \svalue e''\ + SUMM s''''' \ ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts (st''''\accounts := acc\)) (STR ''Victim''))" by simp - moreover have "(case (sender e'', x) of (ad, x) \ \x\) \ (\(ad, x)\{(ad, x). fmlookup s''''' (ad + (STR ''.'' + STR ''balance'')) = Some x}. ReadL\<^sub>i\<^sub>n\<^sub>t x)" - proof (cases rule: member_le_sum[of "(sender e'',x)" "{(ad,x). fmlookup s''''' (ad + (STR ''.'' + STR ''balance'')) = Some x}" "\(ad,x). ReadL\<^sub>i\<^sub>n\<^sub>t x"]) - case 1 - then show ?case using x_def by simp - next - case (2 x) - with III show ?case by auto - next - case 3 - have "inj_on (\(ad, x). (ad + (STR ''.'' + STR ''balance''), x)) {(ad, x). (fmlookup s''''' \ (\ad. ad + (STR ''.'' + STR ''balance''))) ad = Some x}" using balance_inj by simp - then have "finite {(ad, x). (fmlookup s''''' \ (\ad. ad + (STR ''.'' + STR ''balance''))) ad = Some x}" using fmlookup_finite[of "\ad. (ad + (STR ''.'' + STR ''balance''))" s'''''] by simp - then show ?case using finite_subset[of "{(ad, x). fmlookup s''''' (ad + (STR ''.'' + STR ''balance'')) = Some x}" "{(ad, x). fmlookup s''''' (ad + (STR ''.'' + STR ''balance'')) = Some x}"] by auto - qed - then have "ReadL\<^sub>i\<^sub>n\<^sub>t x \ SUMM s'''''" by simp - ultimately have "bal + \svalue e''\ + ReadL\<^sub>i\<^sub>n\<^sub>t x \ ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts (st''''\accounts := acc\)) (STR ''Victim''))" by simp - moreover from ff have "bal\0" by (auto simp add:frame_def) - ultimately have "\svalue e''\ + ReadL\<^sub>i\<^sub>n\<^sub>t x \ ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts (st''''\accounts := acc\)) (STR ''Victim''))" by simp - with abc have "\svalue e''\ + ReadL\<^sub>i\<^sub>n\<^sub>t x < 2^256" by simp - moreover have "fmlookup s''''' (sender e' + (STR ''.'' + STR ''balance'')) = fmlookup s''''' (sender e'' + (STR ''.'' + STR ''balance''))" using send1 by simp - ultimately have "bal + \svalue e''\ \ \accessBalance (accounts (st''''\accounts := acc\)) STR ''Victim''\ - SUMM s'''''" and lck: "fmlookup s''''' (sender e'' + (STR ''.'' + STR ''balance'')) = Some x" and "ReadL\<^sub>i\<^sub>n\<^sub>t x + \svalue e''\ < 2 ^ 256" using ad1 ad2 True II x_def by simp+ - moreover from lck have "(accessStorage (TUInt 256) (sender e'' + (STR ''.'' + STR ''balance'')) s''''') = x" by simp - moreover have "\svalue e''\ + ReadL\<^sub>i\<^sub>n\<^sub>t x \ 0" - proof - - have "ReadL\<^sub>i\<^sub>n\<^sub>t v \ 0" using transfer_val1[OF ******] by simp - with `svalue e' = v` sval1 have "(\svalue e''\::int) \ 0" by simp - moreover from III have "ReadL\<^sub>i\<^sub>n\<^sub>t x \ 0" using x_def by simp - ultimately show ?thesis by simp - qed - ultimately show ?thesis using Read_ShowL_id by simp - qed - moreover have "address e'' = STR ''Victim''" using ad1 ad2 True by simp - ultimately obtain s'''''' where "fmlookup (storage st'''''') (STR ''Victim'') = Some s''''''" and "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st'''''') (STR ''Victim'')) - SUMM s'''''' \ bal" and "POS s''''''" - using deposit_frame[OF * s'''''_def _ yyy] III by auto - then show ?thesis using ff by (auto simp add:frame_def) - qed - next - case False - moreover from f2 have "frame bal (st'''\stack:=emptyStore, memory:=emptyStore\)" using transfer_frame[OF ******] asm by simp - then have "frame bal st''''" and ad1: "address e' = address e''" using f2 0 42(3)[OF a1 ** _ _ _ _ *** _ **** _ _ _ ***** _ _ _ _ e'_def, of "KValue adv" "Value TAddr" TAddr fb "KValue v" "Value t"] asm ******* by (auto simp add:frame_def) - then have f4: "frame bal (st''''\accounts := acc\)" using transfer_frame[OF ******] asm by simp - - moreover from ad1 ad2 have "address e'' \ STR ''Victim'' \ fmlookup e\<^sub>p (STR ''Victim'') = Some (victim, SKIP)" using 0 False by simp - then have "\st6' bal. - frame bal (st''''\accounts := acc\) \ - local.stmt f e\<^sub>p e'' cd' (st''''\accounts := acc\) = Normal ((), st6') \ - frame bal st6'" using 42(4)[OF a1 ** _ _ _ _ *** _ **** _ _ _ ***** _ _ _ _ e'_def _ ******* _ _ _ ******] False asm by (auto simp add:frame_def) - ultimately show ?thesis using ******** by blast - qed - ultimately show "frame bal st'" using a1 * ** *** **** by (auto simp add:frame_def) - qed - qed - qed - qed - next - show "?RHS" (is "?LHS \ ?RHS") - proof - assume "address e = STR ''Victim''" - show ?RHS (is "?A \ (?B \ ?C)") - proof (rule conj3) - show ?A (is "\s val bal x. ?LHS s val bal x") - proof (rule allI[OF allI[OF allI[OF allI]]]) - fix s val bal x - show "?LHS s val bal x" (is "?LHS \ ?RHS") - proof - assume ?LHS - then show ?RHS by simp - qed - qed - next - show ?B (is "\s bal x. ?LHS s bal x") - proof (rule allI[OF allI[OF allI]]) - fix s bal x - show "?LHS s bal x" (is "?LHS \ ?RHS") - proof - assume ?LHS - then show ?RHS by simp - qed - qed - next - show ?C (is "\s bal. ?LHS s bal") - proof (rule allI[OF allI]) - fix s bal - show "?LHS s bal" (is "?LHS \ ?RHS") - proof - assume ?LHS - then show ?RHS by simp - qed - qed - qed - qed - qed - qed - qed - qed -next - case (43 ad ex e\<^sub>p e cd st) - show ?case (is "?LHS \ ?RHS") - proof - assume 0: "fmlookup e\<^sub>p STR ''Victim'' = Some (victim, SKIP)" - show ?RHS (is "\st6'. ?RHS st6'") - proof - fix st' - show "?RHS st'" (is "?LHS \ ?RHS") - proof - assume *: "stmt (TRANSFER ad ex) e\<^sub>p e cd st = Normal ((), st')" - show ?RHS (is "?LHS \ ?RHS") - proof - show "?LHS" - proof - assume asm: "address e \ STR ''Victim''" - show "\bal. frame bal st \ frame bal st'" - proof - fix bal - show "frame bal st \ frame bal st'" - proof - assume ff: "frame bal st" - from * have a1: "(applyf (costs (TRANSFER ad ex) e\<^sub>p e cd) \ (\g. assert Gas (\st. gas st \ g) (modify (\st. st\gas := gas st - g\)))) st = Normal ((), st\gas := gas st - costs (TRANSFER ad ex) e\<^sub>p e cd st\)" by auto - from * obtain v t st'' where **: "expr ex e\<^sub>p e cd (st\gas:=gas st - (costs (TRANSFER ad ex) e\<^sub>p e cd st)\) = Normal ((KValue v, Value t), st'')" - by (auto split:if_split_asm result.split_asm Stackvalue.split_asm Type.split_asm) - from asm ff * ** have f1: "frame bal st''" using 43(1)[OF a1] 0 by (simp add:frame_def) - from * ** obtain adv st''' where ***: "expr ad e\<^sub>p e cd st'' = Normal ((KValue adv, Value TAddr), st''')" - by (auto split:if_split_asm result.split_asm Stackvalue.split_asm Type.split_asm Types.split_asm) - from asm * *** f1 have f2: "frame bal st'''" using asm 43(2)[OF a1 **] 0 by (simp add:frame_def) - from * ** *** obtain acc where *****: "Accounts.transfer (address e) adv v (accounts st''') = Some acc" by (auto split:if_split_asm option.split_asm) - from f2 have f3: "frame bal (st'''\accounts := acc\)" using transfer_frame[OF *****] asm by simp - show "frame bal st'" - proof (cases "fmlookup e\<^sub>p adv") - case None - with a1 * ** *** ***** show ?thesis using f3 by auto - next - case (Some a) - then show ?thesis - proof (cases a) - case (Pair ct f) - moreover define e' where "e' = ffold_init ct (emptyEnv adv (address e) v) (fmdom ct)" - moreover from * ** *** Some Pair ***** e'_def obtain st'''' where ******: "stmt f e\<^sub>p e' emptyStore (st'''\accounts := acc, stack:=emptyStore, memory:=emptyStore\) = Normal ((), st'''')" by (auto split:if_split_asm option.split_asm result.split_asm) - moreover from e'_def have ad: "adv = address e'" using ffold_init_ad_same[of ct "(emptyEnv adv (address e) v)" "(fmdom ct)" e'] by simp - - moreover have f4: "frame bal st''''" - proof (cases "adv = STR ''Victim''") - case True - with 0 ** *** Some Pair have "f = SKIP" using victim_def by simp - with ****** have "st''''= st''' - \accounts := acc, stack := emptyStore, memory := emptyStore, - gas := gas st''' - costs SKIP e\<^sub>p e' emptyStore (st'''\accounts := acc, stack := emptyStore, memory := emptyStore\)\" by (simp split:if_split_asm) - with f3 show ?thesis by (simp add:frame_def) - next - case False - with asm ad have "\st6' bal. - frame bal (st'''\accounts := acc\) \ - local.stmt f e\<^sub>p e' emptyStore (st'''\accounts := acc, stack := emptyStore, memory := emptyStore\) = Normal ((), st6') \ - frame bal st6'" using asm Some Pair 43(3)[OF a1 ** _ _ _ *** _ _ _ _ _ ***** _ _ e'_def, where s'e = "st'''\accounts := acc, stack:=emptyStore, memory:=emptyStore\", of "KValue v" "Value t" t "KValue adv" "Value TAddr" TAddr st''' _ f st''' st''' "()"] using 0 by (simp add:frame_def) - with f3 ****** show ?thesis by blast - qed - moreover from * ** *** Some Pair ***** ****** e'_def have st'_def: "st' = st''''\stack:=stack st''', memory := memory st'''\" by (simp split:if_split_asm) - ultimately show "frame bal st'" apply safe by (simp_all add:frame_def) - qed - qed - qed - qed - qed - next - show "?RHS" (is "?LHS \ ?RHS") - proof - assume ad: "address e = STR ''Victim''" - show ?RHS (is "?A \ ?B \ ?C") - proof (rule conj3) - show ?A (is "\s val bal x. ?LHS s val bal x") - proof (rule allI[OF allI[OF allI[OF allI]]]) - fix s val bal x - show "?LHS s val bal x" (is "?LHS \ ?RHS") - proof - assume ?LHS (is "?A1 \ ?A2 \ ?A3 \ ?A4 \ ?A5") - then have ?A1 and ?A2 and ?A3 and ?A4 and ?A5 by auto - define st'' where "st'' = st\gas := gas st - costs transfer e\<^sub>p e cd st\" - define st''' where "st''' = st''\gas := gas st'' - costs\<^sub>e (LVAL (Id (STR ''bal''))) e\<^sub>p e cd st''\" - define st'''' where "st'''' = st'''\gas := gas st''' - costs\<^sub>e SENDER e\<^sub>p e cd st'''\" - from `?A1` * have c1: "gas st > costs transfer e\<^sub>p e cd st" by (auto split:if_split_asm) - have c2: "gas st'' > costs\<^sub>e (LVAL (Id (STR ''bal''))) e\<^sub>p e cd st''" - proof (rule ccontr) - assume "\ costs\<^sub>e (LVAL (Id (STR ''bal''))) e\<^sub>p e cd st'' < gas st''" - with c1 show False using `?A1` * st''_def st'''_def by auto - qed - with `?A4` `?A5` have 00:"expr (LVAL (Id (STR ''bal''))) e\<^sub>p e cd st'' = Normal ((KValue val, Value (TUInt 256)), st''')" using st'''_def st''_def by simp - moreover have "gas st'''>costs\<^sub>e SENDER e\<^sub>p e cd st'''" - proof (rule ccontr) - assume "\ costs\<^sub>e SENDER e\<^sub>p e cd st''' < gas st'''" - with c1 c2 show False using `?A1` `?A4` `?A5` * st''_def st'''_def by auto - qed - then have **:"expr SENDER e\<^sub>p e cd st''' = Normal ((KValue (sender e), Value TAddr), st'''')" using st''''_def by simp - then obtain acc where ***:"Accounts.transfer (address e) (sender e) val (accounts st'''') = Some acc" - and ****: "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance acc (STR ''Victim'')) = ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st'''') (STR ''Victim'')) - (ReadL\<^sub>i\<^sub>n\<^sub>t val)" - proof - - from `?A1` * c1 00 ** obtain acc where acc_def: "Accounts.transfer (address e) (sender e) val (accounts st'''') = Some acc" using st''''_def st'''_def st''_def by (auto split: option.split_asm) - with ad obtain acc' where *: "subBalance (STR ''Victim'') val (accounts st'''') = Some acc'" - and "addBalance (sender e) val acc' = Some acc" by (simp split: option.split_asm) - moreover from * have "acc' = updateBalance(STR ''Victim'') (ShowL\<^sub>i\<^sub>n\<^sub>t (ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st'''') (STR ''Victim'')) - ReadL\<^sub>i\<^sub>n\<^sub>t val)) (accounts st'''')" by (simp split: if_split_asm) - then have "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance acc' (STR ''Victim'')) = ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st'''') (STR ''Victim'')) - (ReadL\<^sub>i\<^sub>n\<^sub>t val)" using Read_ShowL_id by simp - moreover from `?A5` ad have "sender e \ (STR ''Victim'')" by simp - ultimately have "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance acc (STR ''Victim'')) = ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st'''') (STR ''Victim'')) - (ReadL\<^sub>i\<^sub>n\<^sub>t val)" using addBalance_eq[of "sender e" val acc' acc " STR ''Victim''"] by simp - with acc_def show ?thesis using that by simp - qed - show ?RHS - proof (cases "fmlookup e\<^sub>p (sender e)") - case None - with `?A1` 00 * ** *** have "st' = st''''\accounts := acc\" using c1 st''_def by auto - moreover from `?A2` have "fmlookup (storage st'''') (STR ''Victim'') = Some s" using st''_def st'''_def st''''_def by simp - moreover from `?A2` have "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st'''') (STR ''Victim'')) - (SUMM s + ReadL\<^sub>i\<^sub>n\<^sub>t val) \ bal \ bal \ 0" using st''_def st'''_def st''''_def by simp - with **** have "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance acc (STR ''Victim'')) - SUMM s \ bal \ bal \ 0" by simp - then have "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts (st''''\accounts := acc\)) (STR ''Victim'')) - SUMM s \ bal \ bal \ 0" by simp - ultimately show ?thesis using `?A3` by (simp add:frame_def) - next - case (Some a) - then show ?thesis - proof (cases a) - case (Pair ct f) - moreover define e' where e'_def: "e'=ffold_init ct (emptyEnv (sender e) (address e) val) (fmdom ct)" - ultimately obtain st''''' where *****: "stmt f e\<^sub>p e' emptyStore (st''''\accounts := acc, stack:=emptyStore, memory:=emptyStore\) = Normal ((), st''''')" - and ******: "st' = st'''''\stack:=stack st'''', memory := memory st''''\" using `?A1` 00 ** *** Some * stmt.simps(8)[of SENDER "(LVAL (Id (STR ''bal'')))" e\<^sub>p e cd st] c1 st''_def st'''_def st''''_def by (auto split: result.split_asm unit.split_asm) - from `?A1` * have x1: "(applyf (costs (TRANSFER ad ex) e\<^sub>p e cd) \ (\g. assert Gas - (\st. gas st \ g) - (modify (\st. st\gas := gas st - g\)))) - st = - Normal ((), st'')" using st''_def by (simp split:if_split_asm) - from 00 `?A1` have x2: "expr ex e\<^sub>p e cd st'' = Normal ((KValue val, Value (TUInt 256)), st''')" by simp - have x3: "(KValue val, Value (TUInt 256)) = (KValue val, Value (TUInt 256))" by simp - have x4: "KValue val = KValue val" by simp - have x5: "Value (TUInt 256) = Value (TUInt 256)" by simp - from ** `?A1` have x6: "expr ad e\<^sub>p e cd st''' = Normal ((KValue (sender e), Value TAddr), st'''')" by simp - have x7: "(KValue (sender e), Value TAddr) = (KValue (sender e), Value TAddr)" by simp - have x8: "KValue (sender e) = KValue (sender e)" by simp - have x9: "Value TAddr = Value TAddr" by simp - have x10: "TAddr = TAddr" by simp - have x11: "applyf accounts st'''' = Normal (accounts st'''', st'''')" by simp - from *** have x12: "Accounts.transfer (address e) (sender e) val (accounts st'''') = Some acc" by simp - from Some Pair have x13: "fmlookup e\<^sub>p (sender e) = Some (ct,f)" by simp - have x14: "(ct, f) = (ct, f)" by simp - from e'_def have x15: "e' = ffold_init ct (emptyEnv (sender e) (address e) val) (fmdom ct)" by simp - have x16: "get st'''' = Normal (st'''', st'''')" by simp - have x17: "modify (\st. st\accounts := acc, stack := emptyStore, memory := emptyStore\) st'''' = Normal ((), st''''\accounts := acc, stack := emptyStore, memory := emptyStore\)" by simp - - from `?A2` have "fmlookup (storage st'''') (STR ''Victim'') = Some s" using st''_def st'''_def st''''_def by simp - moreover from `?A2` have "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st'''') (STR ''Victim'')) - (SUMM s + ReadL\<^sub>i\<^sub>n\<^sub>t val) \ bal \ bal \ 0" using st''_def st'''_def st''''_def by simp - with **** have "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance acc (STR ''Victim'')) - SUMM s \ bal \ bal \ 0" by simp - then have "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts (st''''\accounts := acc, stack:=emptyStore, memory:=emptyStore\)) (STR ''Victim'')) - SUMM s \ bal \ bal \ 0" by simp - moreover from `?A5` ad have "sender e \ (STR ''Victim'')" by simp - with e'_def have "address e' \ STR ''Victim''" using ffold_init_ad_same[of ct "(emptyEnv (sender e) (address e) val)" "(fmdom ct)" e'] by simp - ultimately have "frame bal st'''''" using 0 ***** 43(3)[OF x1 x2 x3 x4 x5 x6 x7 x8 x9 x10 x11 x12 x13 x14 x15 x16 x17] `?A3` apply safe by (auto simp add:frame_def) - with "******" show ?RHS by (auto simp add:frame_def) - qed - qed - qed - qed - next - show ?B (is "\s bal x. ?LHS s bal x") - proof (rule allI[OF allI[OF allI]]) - fix s bal x - show "?LHS s bal x" (is "?LHS \ ?RHS") - proof - assume ?LHS - then show ?RHS by simp - qed - qed - next - show ?C (is "\s bal. ?LHS s bal") - proof (rule allI[OF allI]) - fix s bal - show "?LHS s bal" (is "?LHS \ ?RHS") - proof - assume ?LHS - then show ?RHS by simp - qed - qed - qed - qed - qed - qed - qed - qed -next - case (44 id0 tp ex smt e\<^sub>p e\<^sub>v cd st) - show ?case (is "?LHS \ ?RHS") - proof - assume 0: "fmlookup e\<^sub>p STR ''Victim'' = Some (victim, SKIP)" - show ?RHS (is "\st6'. ?RHS st6'") - proof - fix st6' - show "?RHS st6'" (is "?LHS \ ?RHS") - proof - assume *: "stmt (BLOCK ((id0, tp), ex) smt) e\<^sub>p e\<^sub>v cd st = Normal ((), st6')" - show ?RHS (is "?LHS \ ?RHS") - proof - show "?LHS" - proof - assume asm: "address e\<^sub>v \ STR ''Victim''" - show "\bal. frame bal st \ frame bal st6'" - proof - fix bal - show "frame bal st \ frame bal st6'" - proof - assume ff: "frame bal st" - with * have a1: "(applyf (costs(BLOCK ((id0, tp), ex) smt) e\<^sub>p e\<^sub>v cd) \ (\g. assert Gas (\st. gas st \ g) (modify (\st. st\gas := gas st - g\)))) st = Normal ((), st\gas := gas st - costs (BLOCK ((id0, tp), ex) smt) e\<^sub>p e\<^sub>v cd st\)" by auto - from * ff have f1: "frame bal (st\gas := gas st - costs (BLOCK ((id0, tp), ex) smt) e\<^sub>p e\<^sub>v cd st\)" by (simp add:frame_def) - - show "frame bal st6'" - proof (cases ex) - case None - with * obtain cd' e' st' where **:"decl id0 tp None False cd (memory (st\gas := gas st - costs (BLOCK ((id0, tp), ex) smt) e\<^sub>p e\<^sub>v cd st\)) cd e\<^sub>v (st\gas := gas st - costs (BLOCK ((id0, tp), ex) smt) e\<^sub>p e\<^sub>v cd st\) = Normal ((cd', e'), st')" by (auto split:result.split_asm if_split_asm) - with * have f2: "frame bal st'" using decl_frame[OF f1 **] by simp - moreover from * None ** obtain st'' where ***: "stmt smt e\<^sub>p e' cd' st' = Normal ((), st'')" by (simp split:if_split_asm) - moreover from ** have ad: "address e' = address e\<^sub>v" using decl_gas_address by simp - moreover from *** asm f2 ad 0 have "frame bal st''" using 44(3)[OF a1 None _ **, of cd' e'] by (simp add:frame_def) - moreover from * None ** *** have "st6' = st''" by (auto split:if_split_asm) - ultimately show ?thesis using f1 asm by auto - next - case (Some ex') - with * obtain v t st' where **:"expr ex' e\<^sub>p e\<^sub>v cd (st\gas := gas st - costs (BLOCK ((id0, tp), ex) smt) e\<^sub>p e\<^sub>v cd st\) = Normal ((v, t), st')" by (auto split:result.split_asm if_split_asm) - with * f1 Some have f2: "frame bal st'" using 44(1)[OF a1 Some] asm 0 by simp - moreover from Some * ** obtain cd' e' st'' where ***:"decl id0 tp (Some (v,t)) False cd (memory st') cd e\<^sub>v st' = Normal ((cd', e'), st'')" by (auto split:result.split_asm if_split_asm) - with * have f3: "frame bal st''" using decl_frame[OF f2 ***] by simp - moreover from ** *** have ad: "address e' = address e\<^sub>v" using decl_gas_address by simp - moreover from * Some ** *** obtain st''' where ****: "stmt smt e\<^sub>p e' cd' st'' = Normal ((), st''')" by (simp split:if_split_asm) - moreover from **** asm f3 ad have "frame bal st'''" using 44(2)[OF a1 Some ** _ _ ***] 0 by (simp add:frame_def) - moreover from * Some ** *** **** have "st6' = st'''" by (auto split:if_split_asm) - ultimately show ?thesis using f1 asm by auto - qed - qed - qed - qed - next - show "?RHS" (is "?LHS \ ?RHS") - proof - assume ad: "address e\<^sub>v = STR ''Victim''" - show ?RHS (is "?A \ (?B \ ?C)") - proof (rule conj3) - show ?A (is "\s val bal x. ?LHS s val bal x") - proof (rule allI[OF allI[OF allI[OF allI]]]) - fix s val bal x - show "?LHS s val bal x" (is "?LHS \ ?RHS") - proof - assume ?LHS - then show ?RHS by simp - qed - qed - next - show ?B (is "\s bal x. ?LHS s bal x") - proof (rule allI[OF allI[OF allI]]) - fix s bal x - show "?LHS s bal x" (is "?LHS \ ?RHS") - proof - assume ?LHS - then show ?RHS by simp - qed - qed - next - show ?C (is "\s bal. ?LHS s bal") - proof (rule allI[OF allI]) - fix s bal - show "?LHS s bal" (is "?LHS \ ?RHS") - proof - assume ?LHS (is "?A1 \ ?A2 \ ?A3 \ ?A4 \ ?A5") - then have ?A1 and ?A2 and ?A3 and ?A4 and ?A5 by auto - - define st'' where "st'' = st\gas := gas st - costs keep e\<^sub>p e\<^sub>v cd st\" - with `?A2` `?A3` have 00: "fmlookup (storage st'') (STR ''Victim'') = Some s" - and **: "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st'') (STR ''Victim'')) - (SUMM s) \ bal \ bal \ 0 \ POS s" by simp+ - - from `?A1` * st''_def obtain v t st''' cd' e' st'''' st''''' - where ***: "expr mylval e\<^sub>p e\<^sub>v cd st'' = Normal ((v,t), st''')" - and ****: "decl (STR ''bal'') (Value (TUInt 256)) (Some (v, t)) False cd (memory st''') cd e\<^sub>v st''' = Normal ((cd', e'),st'''')" - and *****: "stmt comp e\<^sub>p e' cd' st'''' = Normal ((), st''''')" - and "st6' = st'''''" - by (auto split:if_split_asm result.split_asm) - - obtain s''' where - f1: "fmlookup (storage st''') (STR ''Victim'') = Some s'''" - and v_def: "v = KValue (accessStorage (TUInt 256) (sender e\<^sub>v + (STR ''.'' + STR ''balance'')) s''')" - and t_def: "t = Value (TUInt 256)" - and f2: "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st''') (STR ''Victim'')) - (SUMM s''') \ bal \ bal \ 0 \ POS s'''" - using securelval[OF *** `?A4` 00 ** ad] by auto - - with **** obtain s'''' where - *******: "fmlookup (storage st'''') (STR ''Victim'') = Some s''''" - and bbal: "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st'''') (STR ''Victim'')) - (SUMM s'''') \ bal \ bal \ 0 \ POS s''''" using decl_frame frame_def by auto - - from ad `?A5` have - ad2: "address e' = STR ''Victim''" - and ss: "sender e'\address e'" using decl_gas_address[OF ****] by auto - - then obtain x where - ******: "fmlookup (denvalue e') (STR ''bal'') = Some (Value (TUInt 256), (Stackloc x))" - and lkup: "fmlookup (denvalue e') STR ''balance'' = Some (Storage (STMap TAddr (STValue (TUInt 256))), Storeloc STR ''balance'')" - and "accessStore x (stack st'''') = Some (KValue (accessStorage (TUInt 256) (sender e\<^sub>v + (STR ''.'' + STR ''balance'')) s''''))" - proof - - have "Valuetypes.convert (TUInt 256) (TUInt 256) (accessStorage (TUInt 256) (sender e\<^sub>v + (STR ''.'' + STR ''balance'')) s''') = Some (accessStorage (TUInt 256) (sender e\<^sub>v + (STR ''.'' + STR ''balance'')) s''', TUInt 256)" by simp - with **** v_def t_def have "append (STR ''bal'') (Value (TUInt 256)) (KValue (accessStorage (TUInt 256) (sender e\<^sub>v + (STR ''.'' + STR ''balance'')) s''')) cd e\<^sub>v st''' = Normal ((cd', e'),st'''')" by simp - with f1 v_def t_def have st''''_def: "st'''' = st'''\stack := push v (stack st''')\" and "e' = updateEnv (STR ''bal'') t (Stackloc (ShowL\<^sub>n\<^sub>a\<^sub>t (toploc (stack st''')))) e\<^sub>v" by auto - moreover from ******* f1 st''''_def have "Some (KValue (accessStorage (TUInt 256) (sender e\<^sub>v + (STR ''.'' + STR ''balance'')) s''')) = Some (KValue (accessStorage (TUInt 256) (sender e\<^sub>v + (STR ''.'' + STR ''balance'')) s''''))" by simp - ultimately show ?thesis using t_def v_def `?A4` that by simp - qed - with decl_gas_address **** have sck: "accessStore x (stack st'''') = Some (KValue (accessStorage (TUInt 256) (sender e' + (STR ''.'' + STR ''balance'')) s''''))" by simp - - from * have a1: "(applyf (costs(BLOCK ((id0, tp), ex) smt) e\<^sub>p e\<^sub>v cd) \ (\g. assert Gas (\st. gas st \ g) (modify (\st. st\gas := gas st - g\)))) st = Normal ((), st\gas := gas st - costs (BLOCK ((id0, tp), ex) smt) e\<^sub>p e\<^sub>v cd st\)" by auto - from `?A1` have a2: "ex = Some (LVAL (Ref STR ''balance'' [SENDER]))" by simp - from `?A1` *** have a3: "local.expr (LVAL (Ref STR ''balance'' [SENDER])) e\<^sub>p e\<^sub>v cd - (st\gas := gas st - costs (BLOCK ((id0, tp), ex) smt) e\<^sub>p e\<^sub>v cd st\) = - Normal ((v, t), st''')" using st''_def by simp - from `?A1` **** have a4: "decl id0 tp (Some (v, t)) False cd (memory st''') cd e\<^sub>v st''' = Normal ((cd', e'), st'''')" by simp - from `?A1` ***** `st6' = st'''''` have a5: "local.stmt smt e\<^sub>p e' cd' st'''' = Normal ((), st6')" by simp - show "(\s'. fmlookup (storage st6') STR ''Victim'' = Some s' \ - ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st6') (STR ''Victim'')) - (SUMM s') \ bal \ bal \ 0 \ POS s')" - using 44(2)[OF a1 a2 a3 _ _ a4, of cd' e'] 0 a5 ad2 `?A1` ******* bbal ****** lkup sck `?A4` ss apply safe by auto - qed - qed - qed - qed - qed - qed - qed - qed -qed - -corollary final1: - assumes "fmlookup ep (STR ''Victim'') = Some (victim, SKIP)" - and "stmt (EXTERNAL (ADDRESS (STR ''Victim'')) (STR ''withdraw'') [] val) ep env cd st = Normal((), st')" - and "address env \ (STR ''Victim'')" - and "frame bal st" - shows "frame bal st'" - using assms secure(7)[of ep "(EXTERNAL (ADDRESS (STR ''Victim'')) (STR ''withdraw'') [] val)" env cd st] by simp - -corollary final2: - assumes "fmlookup ep (STR ''Victim'') = Some (victim, SKIP)" - and "stmt (EXTERNAL (ADDRESS (STR ''Victim'')) (STR ''deposit'') [] val) ep env cd st = Normal((), st')" - and "address env \ (STR ''Victim'')" - and "frame bal st" - shows "frame bal st'" - using assms secure(7)[of ep "(EXTERNAL (ADDRESS (STR ''Victim'')) (STR ''deposit'') [] val)" env cd st] by simp - -end - -end +section\Reentrancy\ + +text \ + In the following we use our calculus to verify a contract implementing a simple token. + The contract is defined by definition *bank* and consist of one state variable and two methods: + \<^item> The state variable "balance" is a mapping which assigns a balance to each address. + \<^item> Method "deposit" allows to send money to the contract which is then added to the sender's balance. + \<^item> Method "withdraw" allows to withdraw the callers balance. +\ + +text \ + We then verify that the following invariant (defined by *BANK*) is preserved by both methods: + The difference between + \<^item> the contracts own account-balance and + \<^item> the sum of all the balances kept in the contracts state variable + is larger than a certain threshold. +\ + +text \ + There are two things to note here: + First, Solidity implicitly triggers the call of a so-called fallback method whenever we transfer money to a contract. + In particular if another contract calls "withdraw", this triggers an implicit call to the callee's fallback method. + This functionality was exploited in the infamous DAO attack which we demonstrate it in terms of an example later on. + Since we do not know all potential contracts which call "withdraw", we need to verify our invariant for all possible Solidity programs. +\ + +text \ + The second thing to note is that we were not able to verify that the difference is indeed constant. + During verification it turned out that this is not the case since in the fallback method a contract could just send us additional money without calling "deposit". + In such a case the difference would change. In particular it would grow. However, we were able to verify that the difference does never shrink which is what we actually want to ensure. +\ + +theory Reentrancy + imports Weakest_Precondition Solidity_Evaluator + "HOL-Eisbach.Eisbach_Tools" +begin + +subsection\Example of Re-entrancy\ + +definition "example_env \ + loadProc (STR ''Attacker'') + ([], + ([], SKIP), + ITE (LESS (BALANCE THIS) (UINT 256 125)) + (EXTERNAL (ADDRESS (STR ''BankAddress'')) (STR ''withdraw'') [] (UINT 256 0)) + SKIP) + (loadProc (STR ''Bank'') + ([(STR ''balance'', Var (STMap TAddr (STValue (TUInt 256)))), + (STR ''deposit'', Method ([], True, + ASSIGN + (Ref (STR ''balance'') [SENDER]) + (PLUS (LVAL (Ref (STR ''balance'') [SENDER])) VALUE))), + (STR ''withdraw'', Method ([], True, + ITE (LESS (UINT 256 0) (LVAL (Ref (STR ''balance'') [SENDER]))) + (COMP + (TRANSFER SENDER (LVAL (Ref (STR ''balance'') [SENDER]))) + (ASSIGN (Ref (STR ''balance'') [SENDER]) (UINT 256 0))) + SKIP))], + ([], SKIP), + SKIP) + fmempty)" + +global_interpretation reentrancy: statement_with_gas costs_ex example_env costs_min + defines stmt = "reentrancy.stmt" + and lexp = reentrancy.lexp + and expr = reentrancy.expr + and ssel = reentrancy.ssel + and rexp = reentrancy.rexp + and msel = reentrancy.msel + and load = reentrancy.load + and eval = reentrancy.eval + by unfold_locales auto + +lemma "eval 1000 + (COMP + (EXTERNAL (ADDRESS (STR ''BankAddress'')) (STR ''deposit'') [] (UINT 256 10)) + (EXTERNAL (ADDRESS (STR ''BankAddress'')) (STR ''withdraw'') [] (UINT 256 0))) + (STR ''AttackerAddress'') + (STR ''Attacker'') + (STR '''') + (STR ''0'') + [(STR ''BankAddress'', STR ''100'', Contract (STR ''Bank''), 0), (STR ''AttackerAddress'', STR ''100'', Contract (STR ''Attacker''), 0)] + [] + = STR ''BankAddress: balance==70 - Bank(balance[AttackerAddress]==0\)\AttackerAddress: balance==130 - Attacker()''" + by eval + +subsection\Definition of Contract\ + +abbreviation myrexp::L + where "myrexp \ Ref (STR ''balance'') [SENDER]" + +abbreviation mylval::E + where "mylval \ LVAL myrexp" + +abbreviation assign::S + where "assign \ ASSIGN (Ref (STR ''balance'') [SENDER]) (UINT 256 0)" + +abbreviation transfer::S + where "transfer \ TRANSFER SENDER (LVAL (Id (STR ''bal'')))" + +abbreviation comp::S + where "comp \ COMP assign transfer" + +abbreviation keep::S + where "keep \ BLOCK ((STR ''bal'', Value (TUInt 256)), Some mylval) comp" + +abbreviation deposit::S + where "deposit \ ASSIGN (Ref (STR ''balance'') [SENDER]) (PLUS (LVAL (Ref (STR ''balance'') [SENDER])) VALUE)" + +abbreviation "banklist \ [ + (STR ''balance'', Var (STMap TAddr (STValue (TUInt 256)))), + (STR ''deposit'', Method ([], True, deposit)), + (STR ''withdraw'', Method ([], True, keep))]" + +definition bank::"(Identifier, Member) fmap" + where "bank \ fmap_of_list banklist" + +subsection\Verification\ + +locale Reentrancy = Calculus + + assumes r0: "cname = STR ''Bank''" + and r1: "members = bank" + and r2: "fb = SKIP" + and r3: "const = ([], SKIP)" +begin + +subsubsection\Method lemmas\ +text \ + These lemmas are required by @{term vcg_external}. +\ + +lemma mwithdraw[mcontract]: + "members $$ STR ''withdraw'' = Some (Method ([], True, keep))" + using r1 unfolding bank_def by simp + +lemma mdeposit[mcontract]: + "members $$ STR ''deposit'' = Some (Method ([], True, deposit))" + using r1 unfolding bank_def by simp + +subsubsection\Variable lemma\ + +lemma balance: + "members $$ (STR ''balance'') = Some (Var (STMap TAddr (STValue (TUInt 256))))" + using r1 bank_def fmlookup_of_list[of "[(STR ''balance'', Var (STMap TAddr (STValue (TUInt 256)))), (STR ''deposit'', Method ([], True, ASSIGN myrexp (PLUS mylval VALUE))), + (STR ''withdraw'', Method ([], True, BLOCK ((STR ''bal'', Value (TUInt 256)), Some mylval) (COMP (ASSIGN myrexp (UINT 256 0)) Reentrancy.transfer)))]"] by simp + +subsubsection\Case lemmas\ +text \ + These lemmas are required by @{term vcg_transfer}. +\ +lemma cases_ext: + assumes "members $$ mid = Some (Method (fp,True,f))" + and "fp = [] \ P deposit" + and "fp = [] \ P keep" + shows "P f" +proof - + from assms(1) r1 consider (withdraw) "mid = STR ''withdraw''" | (deposit) "mid = STR ''deposit''" using bank_def fmap_of_list_SomeD[of "[(STR ''balance'', Var (STMap TAddr (STValue (TUInt 256)))), (STR ''deposit'', Method ([], True, deposit)), (STR ''withdraw'', Method ([], True, keep))]"] by auto + then show ?thesis + proof (cases) + case withdraw + then have "f = keep" and "fp = []" using assms(1) bank_def r1 fmlookup_of_list[of banklist] by simp+ + then show ?thesis using assms(3) by simp + next + case deposit + then have "f = deposit" and "fp = []" using assms(1) bank_def r1 fmlookup_of_list[of banklist] by simp+ + then show ?thesis using assms(2) by simp + qed +qed + +lemma cases_int: + assumes "members $$ mid = Some (Method (fp,False,f))" + shows "P fp f" + using assms r1 bank_def fmap_of_list_SomeD[of "[(STR ''balance'', Var (STMap TAddr (STValue (TUInt 256)))), (STR ''deposit'', Method ([], True, deposit)), (STR ''withdraw'', Method ([], True, keep))]"] by auto + +lemma cases_fb: + assumes "P SKIP" + shows "P fb" +using assms bank_def r2 by simp + +lemma cases_cons: + assumes "fst const = [] \ P (fst const, SKIP)" + shows "P const" +using assms bank_def r3 by simp + +subsubsection\Definition of Invariant\ + +abbreviation "SUMM s \ \(ad,x)|fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x. ReadL\<^sub>i\<^sub>n\<^sub>t x" + +abbreviation "POS s \ \ad x. fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x \ ReadL\<^sub>i\<^sub>n\<^sub>t x \ 0" + +definition "iv s a \ a \ SUMM s \ POS s" + +lemma weaken: + assumes "iv (storage st ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc ad)) - ReadL\<^sub>i\<^sub>n\<^sub>t v)" + and "ReadL\<^sub>i\<^sub>n\<^sub>t v \0" + shows "iv (storage st ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc ad)))" +using assms unfolding iv_def by simp + +subsubsection\Additional lemmas\ + +lemma expr_0: + assumes "load True [] xe (ffold (init members) (emptyEnv ad cname (address env) v) (fmdom members)) emptyStore emptyStore emptyStore env cd (st\gas := g1\) g1 = Normal ((e\<^sub>l, cd\<^sub>l, k\<^sub>l, m\<^sub>l), g1')" + and "decl STR ''bal'' (Value (TUInt 256)) (Some (lv, lt)) False cd\<^sub>l m\<^sub>l s (cd\<^sub>l, m\<^sub>l, k\<^sub>l, e\<^sub>l) = Some (cd', mem', sck', e')" + and "expr (UINT 256 0) ev0 cd0 st0 g0 = Normal ((rv, rt),g''a)" + shows "rv= KValue (ShowL\<^sub>i\<^sub>n\<^sub>t 0)" and "rt=Value (TUInt 256)" + using assms(3) by (simp add:expr.simps createUInt_def split:if_split_asm)+ + +lemma load_empty_par: + assumes "load True [] xe (ffold (init members) (emptyEnv ad cname (address env) v) (fmdom members)) emptyStore emptyStore emptyStore env cd (st\gas := g1\) g1 = Normal ((e\<^sub>l, cd\<^sub>l, k\<^sub>l, m\<^sub>l), g1')" + shows "load True [] [] (ffold (init members) (emptyEnv ad cname (address env) v) (fmdom members)) emptyStore emptyStore emptyStore env cd (st\gas := g1\) g1 = Normal ((e\<^sub>l, cd\<^sub>l, k\<^sub>l, m\<^sub>l), g1')" +proof - + have "xe=[]" + proof (rule ccontr) + assume "\ xe=[]" + then obtain x and xa where "xe=x # xa" by (meson list.exhaust) + then have "load True [] xe (ffold (init members) (emptyEnv ad cname (address env) v) (fmdom members)) emptyStore emptyStore emptyStore env cd (st\gas := g1\) g1 = throw Err g1" by (simp add:load.simps) + with assms show False by simp + qed + then show ?thesis using assms(1) by simp +qed + +lemma lexp_myrexp_decl: + assumes "load True [] xe (ffold (init members) (emptyEnv ad cname (address env) v) (fmdom members)) emptyStore emptyStore emptyStore env cd (st\gas := g1\) g1 = Normal ((e\<^sub>l, cd\<^sub>l, k\<^sub>l, m\<^sub>l), g1')" + and "decl STR ''bal'' (Value (TUInt 256)) (Some (lv, lt)) False cd\<^sub>l m\<^sub>l s (cd\<^sub>l, m\<^sub>l, k\<^sub>l, e\<^sub>l) = Some (cd', mem', sck', e')" + and "lexp myrexp e' cd' (st0\accounts := acc, stack := sck', memory := mem', gas := g'a\) g'a = Normal ((rv,rt), g''a)" + shows "rv= LStoreloc (address env + (STR ''.'' + STR ''balance''))" and "rt=Storage (STValue (TUInt 256))" +proof - + have "fmlookup (denvalue (ffold (init members) (emptyEnv ad cname (address env) v) (fmdom members))) (STR ''balance'') = Some (Storage (STMap TAddr (STValue (TUInt 256))), Storeloc (STR ''balance''))" unfolding emptyEnv_def using ffold_init_fmdom balance by simp + moreover have el_def: "e\<^sub>l = (ffold (init members) (emptyEnv ad cname (address env) v) (fmdom members))" using load_empty_par[OF assms(1)] by (simp add:load.simps) + ultimately have "fmlookup (denvalue e\<^sub>l) (STR ''balance'') = Some (Storage (STMap TAddr (STValue (TUInt 256))), Storeloc (STR ''balance''))" using assms by auto + then have *: "fmlookup (denvalue e') (STR ''balance'') = Some (Storage (STMap TAddr (STValue (TUInt 256))), Storeloc (STR ''balance''))" using decl_env[OF assms(2)] by simp + + moreover obtain g'' where "ssel (STMap TAddr (STValue (TUInt 256))) (STR ''balance'') [SENDER] e' cd' (st0\accounts := acc, stack := sck', memory := mem', gas := g'a\) g'a = Normal (((address env) + (STR ''.'' + STR ''balance''), STValue (TUInt 256)), g'')" + proof - + have "g'a > costs\<^sub>e SENDER e' cd' (st0\accounts := acc, stack := sck', memory := mem', gas := g'a\)" using assms(3) * by (simp add:expr.simps ssel.simps lexp.simps split:if_split_asm) + then obtain g'' where "expr SENDER e' cd' (st0\accounts := acc, stack := sck', memory := mem', gas := g'a\) g'a = Normal ((KValue (sender e'), Value TAddr), g'')" by (simp add:expr.simps) + moreover have "sender e\<^sub>l = address env" using el_def ffold_init_ad_same[of members "(emptyEnv ad cname (address env) v)" "fmdom members" e\<^sub>l] unfolding emptyEnv_def by simp + then have "sender e' = address env" using decl_env[OF assms(2)] by simp + ultimately show ?thesis using that hash_def by (simp add:ssel.simps) + qed + ultimately show "rv= LStoreloc (address env + (STR ''.'' + STR ''balance''))" and "rt=Storage (STValue (TUInt 256))" using assms(3) by (simp add:lexp.simps)+ +qed + +lemma expr_bal: + assumes "expr (LVAL (L.Id STR ''bal'')) e' cd' (st\accounts := acc, stack := sck', memory := mem', gas := g''a, storage := (storage st) (address e' := fmupd l v' s'), gas := g''\) g'' = Normal ((KValue lv, Value t), g''')" + and "(sck', e') = astack STR ''bal'' (Value (TUInt 256)) (KValue (accessStorage (TUInt 256) (address env + (STR ''.'' + STR ''balance'')) s')) (k\<^sub>l, e\<^sub>l)" + shows "(\accessStorage (TUInt 256) (address env + (STR ''.'' + STR ''balance'')) s'\::int) = ReadL\<^sub>i\<^sub>n\<^sub>t lv" (is ?G1) and "t = TUInt 256" +proof - + from assms(1) + have "expr (LVAL (L.Id STR ''bal'')) e' cd' (st\accounts := acc, stack := sck', memory := mem', gas := g''a, storage := (storage st) (address e' := fmupd l v' s'), gas := g''\) g'' = rexp ((L.Id STR ''bal'')) e' cd' ((st\accounts := acc, stack := sck', memory := mem', gas := g''a, storage := (storage st) (address e' := fmupd l v' s'), gas := g''\)) (g'' - costs\<^sub>e (LVAL (L.Id STR ''bal'')) e' cd' (st\accounts := acc, stack := sck', memory := mem', gas := g''a, storage := (storage st) (address e' := fmupd l v' s'), gas := g''\))" by (simp add:expr.simps split:if_split_asm) + moreover have "rexp ((L.Id STR ''bal'')) e' cd' ((st\accounts := acc, stack := sck', memory := mem', gas := g''a, storage := (storage st) (address e' := fmupd l v' s'), gas := g''\)) (g'' - costs\<^sub>e (LVAL (L.Id STR ''bal'')) e' cd' (st\accounts := acc, stack := sck', memory := mem', gas := g''a, storage := (storage st) (address e' := fmupd l v' s'), gas := g''\)) = Normal ((KValue (accessStorage (TUInt 256) (address env + (STR ''.'' + STR ''balance'')) s'), (Value (TUInt 256))),(g'' - costs\<^sub>e (LVAL (L.Id STR ''bal'')) e' cd' (st\accounts := acc, stack := sck', memory := mem', gas := g''a, storage := (storage st) (address e' := fmupd l v' s'), gas := g''\)) )" + proof - + from assms(2) have "fmlookup (denvalue e') (STR ''bal'') = Some (Value (TUInt 256), Stackloc \toploc k\<^sub>l\)" by (simp add:push_def allocate_def updateStore_def ) + moreover have "accessStore (\toploc k\<^sub>l\) sck' = Some (KValue (accessStorage (TUInt 256) (address env + (STR ''.'' + STR ''balance'')) s'))" using assms(2) by (simp add:push_def allocate_def updateStore_def accessStore_def) + ultimately show ?thesis by (simp add:rexp.simps) + qed + ultimately have "expr (LVAL (L.Id STR ''bal'')) e' cd' (st\accounts := acc, stack := sck', memory := mem', gas := g''a, storage := (storage st) (address e' := fmupd l v' s'), gas := g''\) g'' = Normal ((KValue (accessStorage (TUInt 256) (address env + (STR ''.'' + STR ''balance'')) s'), (Value (TUInt 256))),(g'' - costs\<^sub>e (LVAL (L.Id STR ''bal'')) e' cd' (st\accounts := acc, stack := sck', memory := mem', gas := g''a, storage := (storage st) (address e' := fmupd l v' s'), gas := g''\)))" and "t = TUInt 256" using assms(1) by simp+ + then have "lv = accessStorage (TUInt 256) (address env + (STR ''.'' + STR ''balance'')) s'" using assms(1) by (auto) + with `t = TUInt 256` show ?G1 and "t = TUInt 256" by simp+ +qed + +lemma lexp_myrexp: + assumes "load True [] xe (ffold (init members) (emptyEnv ad cname (address env) v) (fmdom members)) emptyStore emptyStore emptyStore env cd (st\gas := g1\) g1 = Normal ((e\<^sub>l, cd\<^sub>l, k\<^sub>l, m\<^sub>l), g1')" + and "lexp myrexp e\<^sub>l cd\<^sub>l (st'\gas := g2\) g2 = Normal ((rv,rt), g2')" + shows "rv= LStoreloc (address env + (STR ''.'' + STR ''balance''))" and "rt=Storage (STValue (TUInt 256))" +proof - + have "fmlookup (denvalue (ffold (init members) (emptyEnv ad cname (address env) v) (fmdom members))) (STR ''balance'') = Some (Storage (STMap TAddr (STValue (TUInt 256))), Storeloc (STR ''balance''))" unfolding emptyEnv_def using ffold_init_fmdom balance by simp + moreover have el_def: "e\<^sub>l = (ffold (init members) (emptyEnv ad cname (address env) v) (fmdom members))" using load_empty_par[OF assms(1)] by (simp add:load.simps) + ultimately have *: "fmlookup (denvalue e\<^sub>l) (STR ''balance'') = Some (Storage (STMap TAddr (STValue (TUInt 256))), Storeloc (STR ''balance''))" using assms by auto + + moreover obtain g'' where "ssel (STMap TAddr (STValue (TUInt 256))) (STR ''balance'') [SENDER] e\<^sub>l cd\<^sub>l (st'\gas := g2\) g2 = Normal (((address env) + (STR ''.'' + STR ''balance''), STValue (TUInt 256)), g'')" + proof - + have "g2 > costs\<^sub>e SENDER e\<^sub>l cd\<^sub>l (st'\gas := g2\)" using assms(2) * by (simp add:expr.simps ssel.simps lexp.simps split:if_split_asm) + then obtain g'' where "expr SENDER e\<^sub>l cd\<^sub>l (st'\gas := g2\) g2 = Normal ((KValue (sender e\<^sub>l), Value TAddr), g'')" by (simp add: expr.simps) + moreover have "sender e\<^sub>l = address env" using el_def ffold_init_ad_same[of members "(emptyEnv ad cname (address env) v)" "fmdom members" e\<^sub>l] unfolding emptyEnv_def by simp + ultimately show ?thesis using that hash_def by (simp add:ssel.simps) + qed + ultimately show "rv= LStoreloc (address env + (STR ''.'' + STR ''balance''))" and "rt=Storage (STValue (TUInt 256))" using assms(2) by (simp add: lexp.simps)+ +qed + +lemma expr_balance: + assumes "load True [] xe (ffold (init members) (emptyEnv ad cname (address env) v) (fmdom members)) emptyStore emptyStore emptyStore env cd (st\gas := g1\) g1 = Normal ((e\<^sub>l, cd\<^sub>l, k\<^sub>l, m\<^sub>l), g1')" + and "expr (LVAL (Ref (STR ''balance'') [SENDER])) e\<^sub>l cd\<^sub>l (st\accounts := acc, stack := k\<^sub>l, memory := m\<^sub>l, gas := g2\) g2 = Normal ((va, ta), g'a)" + shows "va= KValue (accessStorage (TUInt 256) (address env + (STR ''.'' + STR ''balance'')) (storage st ad))" + and "ta = Value (TUInt 256)" +proof - + have "fmlookup (denvalue (ffold (init members) (emptyEnv ad cname (address env) v) (fmdom members))) (STR ''balance'') = Some (Storage (STMap TAddr (STValue (TUInt 256))), Storeloc (STR ''balance''))" unfolding emptyEnv_def using ffold_init_fmdom balance by simp + moreover have el_def: "e\<^sub>l = (ffold (init members) (emptyEnv ad cname (address env) v) (fmdom members))" using load_empty_par[OF assms(1)] by (simp add:load.simps) + ultimately have *: "fmlookup (denvalue e\<^sub>l) (STR ''balance'') = Some (Storage (STMap TAddr (STValue (TUInt 256))), Storeloc (STR ''balance''))" using assms by auto + + moreover obtain g'' where "ssel (STMap TAddr (STValue (TUInt 256))) (STR ''balance'') [SENDER] e\<^sub>l cd\<^sub>l (st\accounts := acc, stack := k\<^sub>l, memory := m\<^sub>l, gas := g2\) g2 = Normal (((address env) + (STR ''.'' + STR ''balance''), STValue (TUInt 256)), g'')" + proof - + have "g2 > costs\<^sub>e SENDER e\<^sub>l cd\<^sub>l (st\accounts := acc, stack := k\<^sub>l, memory := m\<^sub>l, gas := g2\)" using assms(2) * by (simp add: expr.simps ssel.simps rexp.simps split:if_split_asm) + then obtain g'' where "expr SENDER e\<^sub>l cd\<^sub>l (st\accounts := acc, stack := k\<^sub>l, memory := m\<^sub>l, gas := g2\) g2 = Normal ((KValue (sender e\<^sub>l), Value TAddr), g'')" by (simp add:expr.simps ssel.simps rexp.simps) + moreover have "sender e\<^sub>l = address env" using el_def ffold_init_ad_same[of members "(emptyEnv ad cname (address env) v)" "fmdom members" e\<^sub>l] unfolding emptyEnv_def by simp + ultimately show ?thesis using that hash_def by (simp add:expr.simps ssel.simps rexp.simps) + qed + moreover have "ad = address e\<^sub>l" using el_def ffold_init_ad_same[of members "(emptyEnv ad cname (address env) v)" "fmdom members" e\<^sub>l] unfolding emptyEnv_def by simp + ultimately show "va= KValue (accessStorage (TUInt 256) (address env + (STR ''.'' + STR ''balance'')) (storage st ad))" and "ta = Value (TUInt 256)" using assms(2) by (simp add:expr.simps ssel.simps rexp.simps split:if_split_asm)+ +qed + +lemma balance_inj: "inj_on (\(ad, x). (ad + (STR ''.'' + STR ''balance''),x)) {(ad, x). (fmlookup y \ f) ad = Some x}" +proof + fix v v' assume asm1: "v \ {(ad, x). (fmlookup y \ f) ad = Some x}" and asm2: "v' \ {(ad, x). (fmlookup y \ f) ad = Some x}" + and *:"(case v of (ad, x) \ (ad + (STR ''.'' + STR ''balance''),x)) = (case v' of (ad, x) \ (ad + (STR ''.'' + STR ''balance''),x))" + then obtain ad ad' x x' where **: "v = (ad,x)" and ***: "v'=(ad',x')" by auto + moreover from * ** *** have "ad + (STR ''.'' + STR ''balance'') = ad' + (STR ''.'' + STR ''balance'')" by simp + with * ** have "ad = ad'" using String_Cancel[of ad "STR ''.'' + STR ''balance''" ad' ] by simp + moreover from asm1 asm2 ** *** have "(fmlookup y \ f) ad = Some x" and "(fmlookup y \ f) ad' = Some x'" by auto + with calculation(3) have "x=x'" by simp + ultimately show "v=v'" by simp +qed + +lemma fmfinite: "finite ({(ad, x). fmlookup y ad = Some x})" +proof - + have "{(ad, x). fmlookup y ad = Some x} \ dom (fmlookup y) \ ran (fmlookup y)" + proof + fix x assume "x \ {(ad, x). fmlookup y ad = Some x}" + then have "fmlookup y (fst x) = Some (snd x)" by auto + then have "fst x \ dom (fmlookup y)" and "snd x \ ran (fmlookup y)" using Map.ranI by (blast,metis) + then show "x \ dom (fmlookup y) \ ran (fmlookup y)" by (simp add: mem_Times_iff) + qed + thus ?thesis by (simp add: finite_ran finite_subset) +qed + +lemma fmlookup_finite: + fixes f :: "'a \ 'a" + and y :: "('a, 'b) fmap" + assumes "inj_on (\(ad, x). (f ad, x)) {(ad, x). (fmlookup y \ f) ad = Some x}" + shows "finite {(ad, x). (fmlookup y \ f) ad = Some x}" +proof (cases rule: inj_on_finite[OF assms(1), of "{(ad, x)|ad x. (fmlookup y) ad = Some x}"]) + case 1 + then show ?case by auto +next + case 2 + then have *: "finite {(ad, x) |ad x. fmlookup y ad = Some x}" using fmfinite[of y] by simp + show ?case using finite_image_set[OF *, of "\(ad, x). (ad, x)"] by simp +qed + +lemma expr_plus: + assumes "load True [] xe (ffold (init members) (emptyEnv ad cname (address env) v) (fmdom members)) emptyStore emptyStore emptyStore env cd (st\gas := g3\) g3 = Normal ((e\<^sub>l, cd\<^sub>l, k\<^sub>l, m\<^sub>l), g3')" + and "expr (PLUS a0 b0) ev0 cd0 st0 g0 = Normal ((xs, t'0), g'0)" + shows "\s. xs = KValue (s)" + using assms by (auto simp add:expr.simps split:if_split_asm result.split_asm prod.split_asm Stackvalue.split_asm Type.split_asm option.split_asm) + +lemma summ_eq_sum: + "SUMM s' = (\(ad,x)|fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ adr. ReadL\<^sub>i\<^sub>n\<^sub>t x) + + ReadL\<^sub>i\<^sub>n\<^sub>t (accessStorage (TUInt 256) (adr + (STR ''.'' + STR ''balance'')) s')" +proof (cases "fmlookup s' (adr + (STR ''.'' + STR ''balance'')) = None") + case True + then have "accessStorage (TUInt 256) (adr + (STR ''.'' + STR ''balance'')) s' = ShowL\<^sub>i\<^sub>n\<^sub>t 0" unfolding accessStorage_def by simp + moreover have "{(ad,x). fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x} = {(ad,x). fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ adr}" + proof + show "{(ad, x). fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x} \ {(ad, x). fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ adr}" + proof + fix x + assume "x \ {(ad, x). fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x}" + then show "x \ {(ad, x). fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ adr}" using True by auto + qed + next + show "{(ad, x). fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ adr} \ {(ad, x). fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x }" + proof + fix x + assume "x \ {(ad, x). fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ adr}" + then show "x \ {(ad, x). fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x}" using True by auto + qed + qed + then have "SUMM s' = (\(ad,x)|fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ adr. ReadL\<^sub>i\<^sub>n\<^sub>t x)" by simp + ultimately show ?thesis using Read_ShowL_id by simp +next + case False + then obtain val where val_def: "fmlookup s' (adr + (STR ''.'' + STR ''balance'')) = Some val" by auto + + have "inj_on (\(ad, x). (ad + (STR ''.'' + STR ''balance''), x)) {(ad, x). (fmlookup s' \ (\ad. ad + (STR ''.'' + STR ''balance''))) ad = Some x}" using balance_inj by simp + then have "finite {(ad, x). (fmlookup s' \ (\ad. ad + (STR ''.'' + STR ''balance''))) ad = Some x}" using fmlookup_finite[of "\ad. (ad + (STR ''.'' + STR ''balance''))" s'] by simp + then have sum1: "finite ({(ad,x). fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ adr})" using finite_subset[of "{(ad, x). fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ adr}" "{(ad, x). fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x}"] by auto + moreover have sum2: "(adr,val) \ {(ad,x). fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ adr}" by simp + moreover from sum1 val_def have "insert (adr,val) {(ad, x). fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ adr} = {(ad, x). fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x}" by auto + ultimately show ?thesis using sum.insert[OF sum1 sum2, of "\(ad,x). ReadL\<^sub>i\<^sub>n\<^sub>t x"] val_def unfolding accessStorage_def by simp +qed + +lemma sum_eq_update: + assumes s''_def: "s'' = fmupd (adr + (STR ''.'' + STR ''balance'')) v' s'" + shows "(\(ad,x)|fmlookup s'' (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ adr. ReadL\<^sub>i\<^sub>n\<^sub>t x) + =(\(ad,x)|fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ adr. ReadL\<^sub>i\<^sub>n\<^sub>t x)" +proof - + have "{(ad,x). fmlookup s'' (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ adr} = {(ad,x). fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ adr}" + proof + show "{(ad, x). fmlookup s'' (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ adr} + \ {(ad, x). fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ adr}" + proof + fix xx + assume "xx \ {(ad, x). fmlookup s'' (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ adr}" + then obtain ad x where "xx = (ad,x)" and "fmlookup s'' (ad + (STR ''.'' + STR ''balance'')) = Some x" and "ad \ adr" by auto + then have "fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x" using String_Cancel[of ad "(STR ''.'' + STR ''balance'')" "adr"] s''_def by (simp split:if_split_asm) + with `ad \ adr` `xx = (ad,x)` show "xx \ {(ad, x). fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ adr}" by simp + qed + next + show "{(ad, x). fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ adr} + \ {(ad, x). fmlookup s'' (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ adr}" + proof + fix xx + assume "xx \ {(ad, x). fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ adr}" + then obtain ad x where "xx = (ad,x)" and "fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x" and "ad \ adr" by auto + then have "fmlookup s'' (ad + (STR ''.'' + STR ''balance'')) = Some x" using String_Cancel[of ad "(STR ''.'' + STR ''balance'')" "adr"] s''_def by (auto split:if_split_asm) + with `ad \ adr` `xx = (ad,x)` show "xx \ {(ad, x). fmlookup s'' (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ adr}" by simp + qed + qed + thus ?thesis by simp +qed + +lemma adapt_deposit: + assumes "address env \ ad" + and "load True [] xe (ffold (init members) (emptyEnv ad cname (address env) v) (fmdom members)) emptyStore emptyStore emptyStore env cd (st0\gas := g3\) g3 = Normal ((e\<^sub>l, cd\<^sub>l, k\<^sub>l, m\<^sub>l), g3')" + and "Accounts.transfer (address env) ad v a = Some acc" + and "iv (storage st0 ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc ad)) - ReadL\<^sub>i\<^sub>n\<^sub>t v)" + and "lexp myrexp e\<^sub>l cd\<^sub>l (st0\gas := g'', accounts := acc, stack := k\<^sub>l, memory := m\<^sub>l, gas := g'\) g' = Normal ((LStoreloc l, Storage (STValue t')), g''a)" + and "expr (PLUS mylval VALUE) e\<^sub>l cd\<^sub>l (st0\gas := g'', accounts := acc, stack := k\<^sub>l, memory := m\<^sub>l, gas := g\) g = Normal ((KValue va, Value ta), g')" + and "Valuetypes.convert ta t' va = Some v'" + shows "(ad = address e\<^sub>l \ iv (storage st0 (address e\<^sub>l)(l $$:= v')) \bal (acc (address e\<^sub>l))\) \ + (ad \ address e\<^sub>l \ iv (storage st0 ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc ad))))" +proof (rule conjI; rule impI) + assume "ad = address e\<^sub>l" + define s' s'' where "s' = storage st0 (address e\<^sub>l)" and "s'' = storage st0 (address e\<^sub>l)(l $$:= v')" + then have "s'' = fmupd l v' s'" by simp + moreover from lexp_myrexp[OF assms(2) assms(5)] have "l = address env + (STR ''.'' + STR ''balance'')" and "t'=TUInt 256" by simp+ + ultimately have s''_s': "s'' = s' (address env + (STR ''.'' + STR ''balance'') $$:= v')" by simp + + have "fmlookup (denvalue (ffold (init members) (emptyEnv ad cname (address env) v) (fmdom members))) (STR ''balance'') = Some (Storage (STMap TAddr (STValue (TUInt 256))), Storeloc (STR ''balance''))" unfolding emptyEnv_def using ffold_init_fmdom balance by simp + moreover have "e\<^sub>l = (ffold (init members) (emptyEnv ad cname (address env) v) (fmdom members))" using load_empty_par[OF assms(2)] by (simp add:load.simps) + ultimately have "fmlookup (denvalue e\<^sub>l) (STR ''balance'') = Some (Storage (STMap TAddr (STValue (TUInt 256))), Storeloc (STR ''balance''))" by simp + then have va_def: "va = (createUInt 256 ((ReadL\<^sub>i\<^sub>n\<^sub>t (accessStorage (TUInt 256) ((sender e\<^sub>l) + (STR ''.'' + STR ''balance'')) s') + ReadL\<^sub>i\<^sub>n\<^sub>t (svalue e\<^sub>l))))" + and "ta = (TUInt 256)" + using assms(6) Read_ShowL_id unfolding s'_def by (auto split:if_split_asm simp add: expr.simps ssel.simps rexp.simps add_def olift.simps hash_def) + + have "svalue e\<^sub>l = v" and "sender e\<^sub>l = address env" and "address e\<^sub>l = ad" using ffold_init_ad_same msel_ssel_expr_load_rexp_gas(4)[OF assms(2)] unfolding emptyEnv_def by simp+ + then have a_frame: "iv s' (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc ad)) - ReadL\<^sub>i\<^sub>n\<^sub>t v)" using assms(4) unfolding s'_def by simp + + from assms(1) have "ReadL\<^sub>i\<^sub>n\<^sub>t (bal (a ad)) + ReadL\<^sub>i\<^sub>n\<^sub>t v < 2^256" using transfer_val2[OF assms(3)] by simp + moreover from `address env \ ad` have "ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc ad)) = ReadL\<^sub>i\<^sub>n\<^sub>t (bal (a ad)) + ReadL\<^sub>i\<^sub>n\<^sub>t v" using transfer_add[OF assms(3)] by simp + ultimately have a_bal: "ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc ad)) < 2^256" by simp + + moreover have a_bounds: + "ReadL\<^sub>i\<^sub>n\<^sub>t (accessStorage (TUInt 256) (sender e\<^sub>l + (STR ''.'' + STR ''balance'')) s') + \svalue e\<^sub>l\ < 2 ^ 256 \ + ReadL\<^sub>i\<^sub>n\<^sub>t (accessStorage (TUInt 256) (sender e\<^sub>l + (STR ''.'' + STR ''balance'')) s') + \svalue e\<^sub>l\ \ 0" + proof (cases "fmlookup s' (sender e\<^sub>l + (STR ''.'' + STR ''balance'')) = None") + case True + hence "(accessStorage (TUInt 256) (sender e\<^sub>l + (STR ''.'' + STR ''balance'')) s') = ShowL\<^sub>i\<^sub>n\<^sub>t 0" unfolding accessStorage_def by simp + moreover have "(\svalue e\<^sub>l\::int) < 2 ^ 256" + proof - + from a_frame have "\svalue e\<^sub>l\ + SUMM s' \ ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc ad))" unfolding iv_def using `svalue e\<^sub>l = v` by simp + moreover have "0 \ SUMM s'" + using a_frame sum_nonneg[of "{(ad,x). fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x}" "\(ad,x). ReadL\<^sub>i\<^sub>n\<^sub>t x"] unfolding iv_def by auto + ultimately have "\svalue e\<^sub>l\ \ ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc ad))" by simp + then show "(\svalue e\<^sub>l\::int) < 2 ^ 256" using a_bal by simp + qed + moreover have "ReadL\<^sub>i\<^sub>n\<^sub>t v \ 0" using transfer_val1[OF assms(3)] by simp + with `svalue e\<^sub>l = v` have "(\svalue e\<^sub>l\::int) \ 0" by simp + ultimately show ?thesis using Read_ShowL_id by simp + next + case False + then obtain x where x_def: "fmlookup s' (sender e\<^sub>l + (STR ''.'' + STR ''balance'')) = Some x" by auto + moreover from a_frame have "\svalue e\<^sub>l\ + SUMM s' \ ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc ad))" unfolding iv_def using `svalue e\<^sub>l = v` by simp + moreover have "(case (sender e\<^sub>l, x) of (ad, x) \ \x\) \ (\(ad, x)\{(ad, x). fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x}. ReadL\<^sub>i\<^sub>n\<^sub>t x)" + proof (cases rule: member_le_sum[of "(sender e\<^sub>l,x)" "{(ad,x). fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x}" "\(ad,x). ReadL\<^sub>i\<^sub>n\<^sub>t x"]) + case 1 + then show ?case using x_def by simp + next + case (2 x) + with a_frame show ?case unfolding iv_def by auto + next + case 3 + have "inj_on (\(ad, x). (ad + (STR ''.'' + STR ''balance''), x)) {(ad, x). (fmlookup s' \ (\ad. ad + (STR ''.'' + STR ''balance''))) ad = Some x}" using balance_inj by simp + then have "finite {(ad, x). (fmlookup s' \ (\ad. ad + (STR ''.'' + STR ''balance''))) ad = Some x}" using fmlookup_finite[of "\ad. (ad + (STR ''.'' + STR ''balance''))" s'] by simp + then show ?case using finite_subset[of "{(ad, x). fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x}" "{(ad, x). fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x}"] by auto + qed + then have "ReadL\<^sub>i\<^sub>n\<^sub>t x \ SUMM s'" by simp + ultimately have "\svalue e\<^sub>l\ + ReadL\<^sub>i\<^sub>n\<^sub>t x \ ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc ad))" by simp + then have "\svalue e\<^sub>l\ + ReadL\<^sub>i\<^sub>n\<^sub>t x \ ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc ad))" by simp + with a_bal have "\svalue e\<^sub>l\ + ReadL\<^sub>i\<^sub>n\<^sub>t x < 2^256" by simp + then have "\svalue e\<^sub>l\ \ \bal (acc ad)\ - SUMM s'" and lck: "fmlookup s' (sender e\<^sub>l + (STR ''.'' + STR ''balance'')) = Some x" and "ReadL\<^sub>i\<^sub>n\<^sub>t x + \svalue e\<^sub>l\ < 2 ^ 256" using a_frame x_def `svalue e\<^sub>l = v` unfolding iv_def by auto + moreover from lck have "(accessStorage (TUInt 256) (sender e\<^sub>l + (STR ''.'' + STR ''balance'')) s') = x" unfolding accessStorage_def by simp + moreover have "\svalue e\<^sub>l\ + ReadL\<^sub>i\<^sub>n\<^sub>t x \ 0" + proof - + have "ReadL\<^sub>i\<^sub>n\<^sub>t v \ 0" using transfer_val1[OF assms(3)] by simp + with `svalue e\<^sub>l = v` have "(\svalue e\<^sub>l\::int) \ 0" by simp + moreover from a_frame have "ReadL\<^sub>i\<^sub>n\<^sub>t x \ 0" unfolding iv_def using x_def by simp + ultimately show ?thesis by simp + qed + ultimately show ?thesis using Read_ShowL_id by simp + qed + ultimately have "va = ShowL\<^sub>i\<^sub>n\<^sub>t (ReadL\<^sub>i\<^sub>n\<^sub>t (accessStorage (TUInt 256) (sender e\<^sub>l + (STR ''.'' + STR ''balance'')) s') + ReadL\<^sub>i\<^sub>n\<^sub>t (svalue e\<^sub>l))" using createUInt_id[where ?v="ReadL\<^sub>i\<^sub>n\<^sub>t (accessStorage (TUInt 256) (sender e\<^sub>l + (STR ''.'' + STR ''balance'')) s') + ReadL\<^sub>i\<^sub>n\<^sub>t (svalue e\<^sub>l)"] va_def by simp + then have v'_def: "v' = ShowL\<^sub>i\<^sub>n\<^sub>t (ReadL\<^sub>i\<^sub>n\<^sub>t (accessStorage (TUInt 256) (address env + (STR ''.'' + STR ''balance'')) s') + ReadL\<^sub>i\<^sub>n\<^sub>t v)" using `sender e\<^sub>l = address env` `svalue e\<^sub>l = v` `t'=TUInt 256` `ta = (TUInt 256)` using assms(7) by auto + + have "SUMM s'' \ \bal (acc ad)\" + proof - + have "SUMM s' \ ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc ad)) - ReadL\<^sub>i\<^sub>n\<^sub>t v" using a_frame unfolding iv_def by simp + moreover have "SUMM s'' = SUMM s' + ReadL\<^sub>i\<^sub>n\<^sub>t v" + proof - + from summ_eq_sum have s1: "SUMM s' = (\(ad,x)|fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ address env. ReadL\<^sub>i\<^sub>n\<^sub>t x) + ReadL\<^sub>i\<^sub>n\<^sub>t (accessStorage (TUInt 256) (address env + (STR ''.'' + STR ''balance'')) s')" by simp + have s2: "SUMM s'' = (\(ad,x)|fmlookup s'' (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ address env. ReadL\<^sub>i\<^sub>n\<^sub>t x) + ReadL\<^sub>i\<^sub>n\<^sub>t (accessStorage (TUInt 256) (address env + (STR ''.'' + STR ''balance'')) s') + ReadL\<^sub>i\<^sub>n\<^sub>t v" + proof - + have "inj_on (\(ad, x). (ad + (STR ''.'' + STR ''balance''), x)) {(ad, x). (fmlookup s'' \ (\ad. ad + (STR ''.'' + STR ''balance''))) ad = Some x}" using balance_inj by simp + then have "finite {(ad, x). (fmlookup s'' \ (\ad. ad + (STR ''.'' + STR ''balance''))) ad = Some x}" using fmlookup_finite[of "\ad. (ad + (STR ''.'' + STR ''balance''))" s''] by simp + then have sum1: "finite ({(ad,x). fmlookup s'' (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ address env})" using finite_subset[of "{(ad, x). fmlookup s'' (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ address env}" "{(ad, x). fmlookup s'' (ad + (STR ''.'' + STR ''balance'')) = Some x}"] by auto + moreover have sum2: "(address env, ShowL\<^sub>i\<^sub>n\<^sub>t (ReadL\<^sub>i\<^sub>n\<^sub>t (accessStorage (TUInt 256) (address env + (STR ''.'' + STR ''balance'')) s') + ReadL\<^sub>i\<^sub>n\<^sub>t v)) \ {(ad,x). fmlookup s'' (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ address env}" by simp + moreover from v'_def s''_s' have "insert (address env, ShowL\<^sub>i\<^sub>n\<^sub>t (ReadL\<^sub>i\<^sub>n\<^sub>t (accessStorage (TUInt 256) (address env + (STR ''.'' + STR ''balance'')) s') + ReadL\<^sub>i\<^sub>n\<^sub>t v)) {(ad, x). fmlookup s'' (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ address env} = {(ad, x). fmlookup s'' (ad + (STR ''.'' + STR ''balance'')) = Some x}" by auto + then show ?thesis using sum.insert[OF sum1 sum2, of "\(ad,x). ReadL\<^sub>i\<^sub>n\<^sub>t x"] Read_ShowL_id by simp + qed + from sum_eq_update[OF s''_s'] have s3: "(\(ad,x)|fmlookup s'' (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ address env. ReadL\<^sub>i\<^sub>n\<^sub>t x) + =(\(ad,x)|fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ address env. ReadL\<^sub>i\<^sub>n\<^sub>t x)" by simp + moreover from s''_s' v'_def have "ReadL\<^sub>i\<^sub>n\<^sub>t (accessStorage (TUInt 256) (address env + (STR ''.'' + STR ''balance'')) s'') = ReadL\<^sub>i\<^sub>n\<^sub>t (accessStorage (TUInt 256) (address env + (STR ''.'' + STR ''balance'')) s') + ReadL\<^sub>i\<^sub>n\<^sub>t v" using Read_ShowL_id unfolding accessStorage_def by simp + ultimately have "SUMM s''= SUMM s' + ReadL\<^sub>i\<^sub>n\<^sub>t v" + proof - + from s2 have "SUMM s'' = (\(ad,x)|fmlookup s'' (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ address env. ReadL\<^sub>i\<^sub>n\<^sub>t x) + ReadL\<^sub>i\<^sub>n\<^sub>t (accessStorage (TUInt 256) (address env + (STR ''.'' + STR ''balance'')) s') + ReadL\<^sub>i\<^sub>n\<^sub>t v" by simp + also from s3 have "\ = (\(ad,x)|fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ address env. ReadL\<^sub>i\<^sub>n\<^sub>t x) + ReadL\<^sub>i\<^sub>n\<^sub>t (accessStorage (TUInt 256) (address env + (STR ''.'' + STR ''balance'')) s') + ReadL\<^sub>i\<^sub>n\<^sub>t v" by simp + also from s1 have "\ = SUMM s' - ReadL\<^sub>i\<^sub>n\<^sub>t (accessStorage (TUInt 256) (address env + (STR ''.'' + STR ''balance'')) s') + ReadL\<^sub>i\<^sub>n\<^sub>t (accessStorage (TUInt 256) (address env + (STR ''.'' + STR ''balance'')) s') + ReadL\<^sub>i\<^sub>n\<^sub>t v" by simp + finally show ?thesis by simp + qed + then show ?thesis by simp + qed + ultimately show ?thesis by simp + qed + moreover have "POS s''" + proof (rule allI[OF allI]) + fix x xa + show "fmlookup s'' (x + (STR ''.'' + STR ''balance'')) = Some xa \ 0 \ (\xa\::int)" + proof + assume a_lookup: "fmlookup s'' (x + (STR ''.'' + STR ''balance'')) = Some xa" + show "0 \ (\xa\::int)" + proof (cases "x = address env") + case True + then show ?thesis using s''_s' a_lookup using Read_ShowL_id v'_def a_bounds `sender e\<^sub>l = address env` `svalue e\<^sub>l = v` by auto + next + case False + then have "fmlookup s' (x + (STR ''.'' + STR ''balance'')) = Some xa" using s''_s' a_lookup String_Cancel[of "address env" "(STR ''.'' + STR ''balance'')" x] by (simp split:if_split_asm) + then show ?thesis using a_frame unfolding iv_def by simp + qed + qed + qed + ultimately show "iv (storage st0 (address e\<^sub>l)(l $$:= v')) \bal (acc (address e\<^sub>l))\" unfolding iv_def s''_def using `ad = address e\<^sub>l` by simp +next + assume "ad \ address e\<^sub>l" + moreover have "ad = address e\<^sub>l" using ffold_init_ad_same msel_ssel_expr_load_rexp_gas(4)[OF assms(2)] unfolding emptyEnv_def by simp + ultimately show "iv (storage st0 ad) \bal (acc ad)\" by simp +qed + +lemma adapt_withdraw: + fixes st acc sck' mem' g''a e' l v' xe + defines "st' \ st\accounts := acc, stack := sck', memory := mem', gas := g''a, storage := (storage st) (address e' := (storage st (address e')) (l $$:= v'))\" + assumes "iv (storage st ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc ad)) - ReadL\<^sub>i\<^sub>n\<^sub>t v)" + and "load True [] xe (ffold (init members) (emptyEnv ad cname (address env) v) (fmdom members)) emptyStore + emptyStore emptyStore env cd (st\gas := g'\) g' = Normal ((e\<^sub>l, cd\<^sub>l, k\<^sub>l, m\<^sub>l), g'')" + and "decl STR ''bal'' (Value (TUInt 256)) (Some (va, ta)) False cd\<^sub>l m\<^sub>l (storage st) (cd\<^sub>l, m\<^sub>l, k\<^sub>l, e\<^sub>l) = + Some (cd', mem', sck', e')" + and "expr (UINT 256 0) e' cd' (st\accounts := acc, stack := sck', memory := mem', gas := ga\) ga = + Normal ((KValue vb, Value tb), g'b)" + and "Valuetypes.convert tb t' vb = Some v'" + and "lexp myrexp e' cd' (st\accounts := acc, stack := sck', memory := mem', gas := g'b\) g'b = + Normal ((LStoreloc l, Storage (STValue t')), g''a)" + and "expr mylval e\<^sub>l cd\<^sub>l (st\accounts := acc, stack := k\<^sub>l, memory := m\<^sub>l, + gas := g'' - costs keep e\<^sub>l cd\<^sub>l (st\gas := g'', accounts := acc, stack := k\<^sub>l, memory := m\<^sub>l\)\) + (g'' - costs keep e\<^sub>l cd\<^sub>l (st\gas := g'', accounts := acc, stack := k\<^sub>l, memory := m\<^sub>l\)) = Normal ((va, ta), g'a)" + and "Accounts.transfer (address env) ad v (accounts st) = Some acc" + and "expr SENDER e' cd' (st' \gas := g\) g = Normal ((KValue adv, Value TAddr), g'x)" + and adv_def: "adv \ ad" + and bal: "expr (LVAL (L.Id STR ''bal'')) e' cd' (st'\gas := g''b\) g''b = Normal ((KValue lv, Value t), g''')" + and con: "Valuetypes.convert t (TUInt 256) lv = Some lv'" + shows "iv (storage st' ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st' ad)) - (ReadL\<^sub>i\<^sub>n\<^sub>t lv'))" +proof - + define acca where "acca = accounts st' ad" + define s' where "s' = storage st (address e')" + define s'a where "s'a = storage st' ad" + moreover have "address e' = ad" + proof - + have "address e' = address e\<^sub>l" using decl_env[OF assms(4)] by simp + also have "address e\<^sub>l = address (ffold (init members) (emptyEnv ad cname (address env) v) (fmdom members))" using msel_ssel_expr_load_rexp_gas(4)[OF assms(3)] by simp + also have "\ = ad" unfolding emptyEnv_def using ffold_init_ad_same by simp + finally show ?thesis . + qed + ultimately have s'a_def: "s'a = fmupd l v' s'" unfolding s'_def st'_def by simp + + have "sender e' = address env" + proof - + have "sender e' = sender e\<^sub>l" using decl_env[OF assms(4)] by simp + also have "sender e\<^sub>l = sender (ffold (init members) (emptyEnv ad cname (address env) v) (fmdom members))" using msel_ssel_expr_load_rexp_gas(4)[OF assms(3)] by simp + also have "\ = address env" unfolding emptyEnv_def using ffold_init_ad_same by simp + finally show ?thesis . + qed + + have "iv s'a (\bal acca\ - \lv'\)" unfolding iv_def + proof intros + have "SUMM s' \ \bal acca\" + proof - + from `address e' = ad` have "iv s' (ReadL\<^sub>i\<^sub>n\<^sub>t (bal acca) - ReadL\<^sub>i\<^sub>n\<^sub>t v)" using assms(2,5) unfolding s'_def st'_def acca_def by simp + then have "SUMM s' \ \bal (acca)\ - \v\" unfolding iv_def by simp + moreover have "ReadL\<^sub>i\<^sub>n\<^sub>t v \ 0" using transfer_val1[OF assms(9)] by simp + ultimately show ?thesis by simp + qed + moreover have "SUMM s'a = SUMM s' - ReadL\<^sub>i\<^sub>n\<^sub>t lv'" + proof - + from summ_eq_sum have "SUMM s'a = (\(ad,x)|fmlookup s'a (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender e'. ReadL\<^sub>i\<^sub>n\<^sub>t x) + ReadL\<^sub>i\<^sub>n\<^sub>t (accessStorage (TUInt 256) (sender e' + (STR ''.'' + STR ''balance'')) s'a)" by simp + moreover from summ_eq_sum have "SUMM s' = (\(ad,x)|fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender e'. ReadL\<^sub>i\<^sub>n\<^sub>t x) + ReadL\<^sub>i\<^sub>n\<^sub>t (accessStorage (TUInt 256) (sender e' + (STR ''.'' + STR ''balance'')) s')" by simp + moreover from s'a_def lexp_myrexp_decl(1)[OF assms(3,4) assms(7)] have " s'a = s'((sender e' + (STR ''.'' + STR ''balance''))$$:= v')" using `sender e' = address env` by simp + with sum_eq_update have "(\(ad,x)|fmlookup s'a (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender e'. ReadL\<^sub>i\<^sub>n\<^sub>t x) = (\(ad,x)|fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender e'. ReadL\<^sub>i\<^sub>n\<^sub>t x)" by simp + moreover have "ReadL\<^sub>i\<^sub>n\<^sub>t (accessStorage (TUInt 256) (sender e' + (STR ''.'' + STR ''balance'')) s'a) = ReadL\<^sub>i\<^sub>n\<^sub>t (accessStorage (TUInt 256) (sender e' + (STR ''.'' + STR ''balance'')) s') - ReadL\<^sub>i\<^sub>n\<^sub>t lv'" + proof - + have "ReadL\<^sub>i\<^sub>n\<^sub>t (accessStorage (TUInt 256) (sender e' + (STR ''.'' + STR ''balance'')) s') = ReadL\<^sub>i\<^sub>n\<^sub>t lv'" + proof - + from expr_balance[OF assms(3) assms(8)] have "va= KValue (accessStorage (TUInt 256) (address env + (STR ''.'' + STR ''balance'')) s')" and "ta = Value (TUInt 256)" using `address e' = ad` unfolding s'_def by simp+ + then have "(sck',e') = astack STR ''bal'' (Value (TUInt 256)) (KValue (accessStorage (TUInt 256) (address env + (STR ''.'' + STR ''balance'')) s')) (k\<^sub>l, e\<^sub>l)" using decl.simps(2) assms(4) by simp + then have "ReadL\<^sub>i\<^sub>n\<^sub>t (accessStorage (TUInt 256) (address env + (STR ''.'' + STR ''balance'')) s') = ReadL\<^sub>i\<^sub>n\<^sub>t lv" and "t = TUInt 256" using expr_bal bal unfolding s'_def st'_def by auto + moreover from `t = TUInt 256` have "lv = lv'" using con by simp + ultimately show ?thesis using `sender e' = address env` by simp + qed + moreover have "ReadL\<^sub>i\<^sub>n\<^sub>t (accessStorage (TUInt 256) (sender e' + (STR ''.'' + STR ''balance'')) s'a) = ReadL\<^sub>i\<^sub>n\<^sub>t (ShowL\<^sub>i\<^sub>n\<^sub>t 0)" + proof - + have "v' = ShowL\<^sub>i\<^sub>n\<^sub>t 0" + proof - + have "vb = createUInt 256 0" and "tb=TUInt 256" using assms(5) by (simp add: expr.simps load.simps split:if_split_asm)+ + moreover have "t'=TUInt 256" using lexp_myrexp_decl(2)[OF assms(3,4) assms(7)] by simp + ultimately show ?thesis using assms(6) by (simp add: createUInt_id) + qed + moreover have "l= (sender e' + (STR ''.'' + STR ''balance''))" using lexp_myrexp_decl(1)[OF assms(3,4) assms(7)] `sender e' = address env` by simp + ultimately show ?thesis using s'a_def accessStorage_def by simp + qed + ultimately show ?thesis using Read_ShowL_id by simp + qed + ultimately show ?thesis by simp + qed + ultimately show "SUMM s'a \ \bal acca\ - \lv'\" by simp + next + fix ad' x + assume *: "fmlookup s'a (ad' + (STR ''.'' + STR ''balance'')) = Some x" + show "0 \ ReadL\<^sub>i\<^sub>n\<^sub>t x" + proof (cases "ad' = sender e'") + case True + moreover have "v' = ShowL\<^sub>i\<^sub>n\<^sub>t 0" + proof - + have "vb = createUInt 256 0" and "tb=TUInt 256" using assms(5) by (simp add:expr.simps split:if_split_asm)+ + moreover have "t'=TUInt 256" using lexp_myrexp_decl(2)[OF assms(3,4) assms(7)] by simp + ultimately show ?thesis using assms(6) by (simp add: createUInt_id) + qed + moreover have "l= (sender e' + (STR ''.'' + STR ''balance''))" using lexp_myrexp_decl(1)[OF assms(3,4) assms(7)] `sender e' = address env` by simp + ultimately show ?thesis using Read_ShowL_id s'a_def * by auto + next + case False + moreover from `address e' = ad` have "iv s' (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc ad)) - ReadL\<^sub>i\<^sub>n\<^sub>t v)" using assms(2) unfolding s'_def by simp + then have "POS s'" unfolding iv_def by simp + moreover have "l= (sender e' + (STR ''.'' + STR ''balance''))" using lexp_myrexp_decl(1)[OF assms(3,4) assms(7)] `sender e' = address env` by simp + ultimately show ?thesis using s'a_def * String_Cancel by (auto split:if_split_asm) + qed + qed + then show ?thesis unfolding s'a_def s'_def acca_def st'_def `address e' = ad` by simp +qed + +lemma wp_deposit[external]: + assumes "address ev \ ad" + and "expr adex ev cd (st\gas := gas st0 - costs (EXTERNAL adex mid xe val) ev cd st0\) (gas st0 - costs (EXTERNAL adex mid xe val) ev cd st0) = Normal ((KValue ad, Value TAddr), g)" + and "expr val ev cd (st0\gas := g\) g = Normal ((KValue v, Value t), g')" + and "Valuetypes.convert t (TUInt 256) v = Some v'" + and "load True [] xe (ffold (init members) (emptyEnv ad cname (address ev) v') (fmdom members)) emptyStore emptyStore emptyStore ev cd (st0\gas := g'\) g' = Normal ((e\<^sub>l, cd\<^sub>l, k\<^sub>l, m\<^sub>l), g'')" + and "Accounts.transfer (address ev) ad v' (accounts st0) = Some acc" + and "iv (storage st0 ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc ad)) - ReadL\<^sub>i\<^sub>n\<^sub>t v')" + shows "wpS (stmt (ASSIGN myrexp (PLUS mylval VALUE)) e\<^sub>l cd\<^sub>l) + (\st. (iv (storage st ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st ad))))) (\e. e = Gas \ e = Err) + (st0\gas := g'', accounts := acc, stack := k\<^sub>l, memory := m\<^sub>l\)" + apply (vcg_assign expr_rule: expr_plus[OF assms(5)] lexp_rule: lexp_myrexp(1)[OF assms(5)]) + by (simp add: adapt_deposit[OF assms(1,5,6,7)]) + +lemma wptransfer: + fixes st0 acc sck' mem' g''a e' l v' + defines "st' \ st0\accounts := acc, stack := sck', memory := mem', gas := g''a, + storage := (storage st0)(address e' := storage st0 (address e')(l $$:= v'))\" + assumes "Pfe ad iv st'" + and "address ev \ ad" + and "g'' \ gas st" + and "type (acc ad) = Some (Contract cname)" + and "expr adex ev cd (st0\gas := gas st0 - costs (EXTERNAL adex mid xe val) ev cd st0\) + (gas st0 - costs (EXTERNAL adex mid xe val) ev cd st0) = + Normal ((KValue ad, Value TAddr), g)" + and "expr val ev cd (st0\gas := g\) g = Normal ((KValue gv, Value gt), g')" + and "Valuetypes.convert gt (TUInt 256) gv = Some gv'" + and "load True [] xe (ffold (init members) (emptyEnv ad cname (address ev) gv') (fmdom members)) emptyStore + emptyStore emptyStore ev cd (st0\gas := g'\) g' = + Normal ((e\<^sub>l, cd\<^sub>l, k\<^sub>l, m\<^sub>l), g'')" + and "Accounts.transfer (address ev) ad gv' (accounts st0) = Some acc" + and "iv (storage st0 ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc ad)) - ReadL\<^sub>i\<^sub>n\<^sub>t gv')" + and "decl STR ''bal'' (Value (TUInt 256)) (Some (v, t)) False cd\<^sub>l m\<^sub>l (storage st0) + (cd\<^sub>l, m\<^sub>l, k\<^sub>l, e\<^sub>l) = Some (cd', mem', sck', e')" + and "Valuetypes.convert ta t' va = Some v'" + and "lexp myrexp e' cd' (st0\accounts := acc, stack := sck', memory := mem', gas := g'a\) g'a = + Normal ((LStoreloc l, Storage (STValue t')), g''a)" + and "expr mylval e\<^sub>l cd\<^sub>l (st0\accounts := acc, stack := k\<^sub>l, memory := m\<^sub>l, + gas := g'' - costs (BLOCK ((STR ''bal'', Value (TUInt 256)), Some mylval) + (COMP (ASSIGN myrexp (UINT 256 0)) Reentrancy.transfer)) e\<^sub>l cd\<^sub>l (st0\gas := g'', accounts := acc, stack := k\<^sub>l, memory := m\<^sub>l\)\) + (g'' - costs (BLOCK ((STR ''bal'', Value (TUInt 256)), Some mylval) (COMP (ASSIGN myrexp (UINT 256 0)) Reentrancy.transfer)) e\<^sub>l + cd\<^sub>l (st0\gas := g'', accounts := acc, stack := k\<^sub>l, memory := m\<^sub>l\)) = + Normal ((v, t), g''')" + and "expr (UINT 256 0) e' cd' (st0\accounts := acc, stack := sck', memory := mem', gas := ga\) ga = Normal ((KValue va, Value ta), g'a)" + shows "wpS (stmt Reentrancy.transfer e' cd') (\st. iv (storage st ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st ad)))) + (\e. e = Gas \ e = Err) st'" +proof (rule meta_eq_to_obj_eq[THEN iffD1, OF Pfe_def assms(2),rule_format], (rule conjI; (rule conjI)?)) + show "address e' = ad" + proof - + have "address e' = address e\<^sub>l" using decl_env[OF assms(12)] by simp + also have "address e\<^sub>l = address (ffold (init members) (emptyEnv ad cname (address ev) gv') (fmdom members))" using msel_ssel_expr_load_rexp_gas(4)[OF assms(9)] by simp + also have "\ = ad" unfolding emptyEnv_def using ffold_init_ad_same by simp + finally show ?thesis . + qed +next + show "\adv g. expr SENDER e' cd' (st'\gas := gas st' - costs Reentrancy.transfer e' cd' st'\) (gas st' - costs Reentrancy.transfer e' cd' st') = Normal ((KValue adv, Value TAddr), g) + \ adv \ ad" + proof (intros) + fix adv g + assume "expr SENDER e' cd' (st'\gas := gas st' - costs Reentrancy.transfer e' cd' st'\) (gas st' - costs Reentrancy.transfer e' cd' st') = Normal ((KValue adv, Value TAddr), g)" + moreover have "sender e' \ ad" + proof - + have "sender e' = sender e\<^sub>l" using decl_env[OF assms(12)] by simp + also have "sender e\<^sub>l = sender (ffold (init members) (emptyEnv ad cname (address ev) gv') (fmdom members))" using msel_ssel_expr_load_rexp_gas(4)[OF assms(9)] by simp + also have "\ = address ev" unfolding emptyEnv_def using ffold_init_ad_same by simp + finally show ?thesis using assms(3) by simp + qed + ultimately show "adv \ ad" by (simp add:expr.simps split:result.split_asm if_split_asm prod.split_asm) + qed +next + show "\adv g v t g' v'. + local.expr SENDER e' cd' (st'\gas := gas st' - costs Reentrancy.transfer e' cd' st'\) + (gas st' - costs Reentrancy.transfer e' cd' st') = Normal ((KValue adv, Value TAddr), g) \ + adv \ ad \ + local.expr (LVAL (L.Id STR ''bal'')) e' cd' (st'\gas := g\) g = Normal ((KValue v, Value t), g') \ + Valuetypes.convert t (TUInt 256) v = Some v' \ + iv (storage st' ad) (\bal (accounts st' ad)\ - ReadL\<^sub>i\<^sub>n\<^sub>t v')" + proof elims + fix adv lg lv lt lg' lv' + assume a1:"expr SENDER e' cd' (st'\gas := gas st' - costs Reentrancy.transfer e' cd' st'\) (gas st' - costs Reentrancy.transfer e' cd' st') = + Normal ((KValue adv, Value TAddr), lg)" + and adv_def: "adv \ ad" + and bal: "expr (LVAL (L.Id STR ''bal'')) e' cd' (st'\gas := lg\) lg = Normal ((KValue lv, Value lt), lg')" + and con: "Valuetypes.convert lt (TUInt 256) lv = Some lv'" + show "iv (storage st' ad) (\bal (accounts st' ad)\ - ReadL\<^sub>i\<^sub>n\<^sub>t lv')" + using adapt_withdraw[where ?acc=acc, OF assms(11) load_empty_par[OF assms(9)] assms(12,16,13,14,15,10) _ adv_def _] a1 bal con unfolding st'_def by simp + qed +qed + +lemma wp_withdraw[external]: + assumes "\st'. gas st' \ gas st \ type (accounts st' ad) = Some (Contract cname) \ Pe ad iv st' \ Pi ad (\_ _. True) (\_ _. True) st' \ Pfi ad (\_. True) (\_. True) st' \ Pfe ad iv st'" + and "address ev \ ad" + and "g'' \ gas st" + and "type (acc ad) = Some (Contract cname)" + and "expr adex ev cd (st0\gas := gas st0 - costs (EXTERNAL adex mid xe val) ev cd st0\) + (gas st0 - costs (EXTERNAL adex mid xe val) ev cd st0) = Normal ((KValue ad, Value TAddr), g)" + and "expr val ev cd (st0\gas := g\) g = Normal ((KValue v, Value t), g')" + and "Valuetypes.convert t (TUInt 256) v = Some v'" + and "load True [] xe (ffold (init members) (emptyEnv ad cname (address ev) v') (fmdom members)) + emptyStore emptyStore emptyStore ev cd (st0\gas := g'\) g' = Normal ((e\<^sub>l, cd\<^sub>l, k\<^sub>l, m\<^sub>l), g'')" + and "Accounts.transfer (address ev) ad v' (accounts st0) = Some acc" + and "iv (storage st0 ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc ad)) - ReadL\<^sub>i\<^sub>n\<^sub>t v')" + shows "wpS (stmt keep e\<^sub>l cd\<^sub>l) + (\st. iv (storage st ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st ad)))) (\e. e = Gas \ e = Err) + (st0\gas := g'', accounts := acc, stack := k\<^sub>l, memory := m\<^sub>l\)" + apply vcg_block_some + apply vcg_comp + apply (vcg_assign expr_rule: expr_0[OF assms(8)] lexp_rule: lexp_myrexp_decl[OF assms(8)]) + apply (rule wptransfer[OF _ assms(2-10)]) + apply (rule_tac assms(1)[THEN conjunct2,THEN conjunct2,THEN conjunct2]) + using assms(3,4) by simp + +lemma wp_fallback: + assumes "Accounts.transfer (address ev) ad v (accounts st0) = Some acca" + and "iv (storage st0 ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acca ad)) - ReadL\<^sub>i\<^sub>n\<^sub>t v)" + shows "wpS (stmt SKIP (ffold (init members) (emptyEnv ad cname (address ev) vc) (fmdom members)) emptyStore) + (\st. iv (storage st ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st ad)))) (\e. e = Gas \ e = Err) + (st0\gas := g'c, accounts := acca, stack := emptyStore, memory := emptyStore\)" + apply vcg_skip + using weaken[where ?acc=acca, OF assms(2) transfer_val1[OF assms(1)]] by auto + +lemma wp_construct: + assumes "Accounts.transfer (address ev) (hash (address ev) \contracts (accounts st (address ev))\) v + ((accounts st) (hash (address ev) \contracts (accounts st (address ev))\ := \bal = ShowL\<^sub>i\<^sub>n\<^sub>t 0, type = Some (Contract i), contracts = 0\)) = + Some acc" + shows "iv fmempty \bal (acc (hash (address ev) \contracts (accounts st (address ev))\))\" +proof - + define adv where "adv = (hash (address ev) \contracts (accounts st (address ev))\)" + moreover have "ReadL\<^sub>i\<^sub>n\<^sub>t (bal (((accounts st) (adv := \bal = ShowL\<^sub>i\<^sub>n\<^sub>t 0, type = Some (Contract i), contracts = 0\)) adv)) = 0" using Read_ShowL_id[of 0] by simp + ultimately have "ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc (hash (address ev) \contracts (accounts st (address ev))\))) \ 0" using transfer_mono[OF assms(1)] by simp + then show ?thesis unfolding iv_def by simp +qed + +lemma wp_true[external]: + assumes "E Gas" + and "E Err" + shows "wpS (stmt f e cd) (\st. True) E st" + unfolding wpS_def wp_def +proof ((split result.split, (split prod.split)?); rule conjI; (rule allI)+; (rule impI)+) + fix x1 x1a s + assume "x1 = (x1a, s)" and "stmt f e cd st = Normal x1" + then show "gas s \ gas st \ True" using stmt_gas by simp +next + fix x2 + assume "stmt f e cd st = Exception x2" + show "E x2" using assms Ex.nchotomy by metis +qed + +subsubsection\Final results\ + +interpretation vcg:VCG costs\<^sub>e ep costs cname members const fb "\_. True" "\_. True" "\_ _. True" "\_ _. True" +by (simp add: VCG.intro Calculus_axioms) + +lemma safe_external: + "external ad iv" + apply (external cases: cases_ext) + apply (rule wp_deposit;assumption) + apply vcg_block_some + apply vcg_comp + apply (vcg_assign expr_rule: expr_0 lexp_rule: lexp_myrexp_decl) + apply (vcg.vcg_transfer_ext ad fallback_int: wp_true fallback_ext: wp_fallback cases_ext:cases_ext cases_int:cases_fb cases_fb:cases_int invariant:adapt_withdraw) + done + +lemma safe_fallback: "fallback ad iv" + apply (fallback cases: cases_fb) + apply (rule wp_fallback; assumption) + done + +lemma safe_constructor: "constructor iv" + apply (constructor cases: cases_cons) + apply vcg_skip + apply (simp add:wp_construct) + done + +theorem safe: + assumes "address ev \ ad" + and "type (accounts st ad) = Some (Contract cname) \ iv (storage st ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st ad)))" + shows "\(st'::State). stmt f ev cd st = Normal ((), st') \ type (accounts st' ad) = Some (Contract cname) + \ iv (storage st' ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st' ad)))" + apply (rule sound) using assms safe_external safe_fallback safe_constructor by auto + +end + +end diff --git a/thys/Solidity/Solidity_Evaluator.thy b/thys/Solidity/Solidity_Evaluator.thy --- a/thys/Solidity/Solidity_Evaluator.thy +++ b/thys/Solidity/Solidity_Evaluator.thy @@ -1,624 +1,694 @@ -section\ Solidty Evaluator and Code Generator Setup\ - -theory - Solidity_Evaluator -imports - Solidity_Main - "HOL-Library.Code_Target_Numeral" - "HOL-Library.Sublist" - "HOL-Library.Finite_Map" -begin - -subsection\Code Generator Setup and Local Tests\ - - -subsubsection\Utils\ - -definition FAILURE::"String.literal" where "FAILURE = STR ''Failure''" -definition "inta_of_int = int o nat_of_integer" -definition "nat_of_int = nat_of_integer" - -fun astore :: "Identifier \ Type \ Valuetype \ StorageT * Environment \ StorageT * Environment" - where "astore i t v (s, e) = (fmupd i v s, (updateEnv i t (Storeloc i) e))" - -subsubsection\Valuetypes\ - -fun dump\<^sub>V\<^sub>a\<^sub>l\<^sub>u\<^sub>e\<^sub>t\<^sub>y\<^sub>p\<^sub>e\<^sub>s::"Types \ Valuetype \ String.literal" where - "dump\<^sub>V\<^sub>a\<^sub>l\<^sub>u\<^sub>e\<^sub>t\<^sub>y\<^sub>p\<^sub>e\<^sub>s (TSInt _) n = n" - | "dump\<^sub>V\<^sub>a\<^sub>l\<^sub>u\<^sub>e\<^sub>t\<^sub>y\<^sub>p\<^sub>e\<^sub>s (TUInt _) n = n" - | "dump\<^sub>V\<^sub>a\<^sub>l\<^sub>u\<^sub>e\<^sub>t\<^sub>y\<^sub>p\<^sub>e\<^sub>s TBool b = (if b = (STR ''True'') then STR ''true'' else STR ''false'')" - | "dump\<^sub>V\<^sub>a\<^sub>l\<^sub>u\<^sub>e\<^sub>t\<^sub>y\<^sub>p\<^sub>e\<^sub>s TAddr ad = ad" - -paragraph\Generalized Unit Tests\ - -lemma "createSInt 8 500 = STR ''-12''" - by(eval) - -lemma "STR ''-92134039538802366542421159375273829975'' - = createSInt 128 45648483135649456465465452123894894554654654654654646999465" - by(eval) - -lemma "STR ''-128'' = createSInt 8 (-128)" - by(eval) - -lemma "STR ''244'' = (createUInt 8 500)" - by(eval) - -lemma "STR ''220443428915524155977936330922349307608'' - = (createUInt 128 4564848313564945646546545212389489455465465465465464699946544654654654654168)" - by(eval) - -lemma "less (TUInt 144) (TSInt 160) (STR ''5'') (STR ''8'') = Some(STR ''True'', TBool) " - by(eval) - - -subsubsection\Load: Accounts\ - -fun load\<^sub>A\<^sub>c\<^sub>c\<^sub>o\<^sub>u\<^sub>n\<^sub>t\<^sub>s :: "Accounts \ (Address \ Balance) list \ Accounts" where - "load\<^sub>A\<^sub>c\<^sub>c\<^sub>o\<^sub>u\<^sub>n\<^sub>t\<^sub>s acc [] = acc" - | "load\<^sub>A\<^sub>c\<^sub>c\<^sub>o\<^sub>u\<^sub>n\<^sub>t\<^sub>s acc ((ad, bal)#as) = load\<^sub>A\<^sub>c\<^sub>c\<^sub>o\<^sub>u\<^sub>n\<^sub>t\<^sub>s (updateBalance ad bal acc) as" - -definition dump\<^sub>A\<^sub>c\<^sub>c\<^sub>o\<^sub>u\<^sub>n\<^sub>t\<^sub>s :: "Accounts \ Address list \ String.literal" -where - "dump\<^sub>A\<^sub>c\<^sub>c\<^sub>o\<^sub>u\<^sub>n\<^sub>t\<^sub>s acc = foldl (\ t a . String.implode ( (String.explode t) - @ (String.explode a) - @ ''.balance'' - @ ''=='' - @ String.explode (accessBalance acc a) - @ ''\'')) - (STR '''')" - -definition init\<^sub>A\<^sub>c\<^sub>c\<^sub>o\<^sub>u\<^sub>n\<^sub>t::"(Address \ Balance) list => Accounts" where - "init\<^sub>A\<^sub>c\<^sub>c\<^sub>o\<^sub>u\<^sub>n\<^sub>t = load\<^sub>A\<^sub>c\<^sub>c\<^sub>o\<^sub>u\<^sub>n\<^sub>t\<^sub>s emptyAccount" - - -subsubsection\Load: Store\ - -type_synonym Data\<^sub>S\<^sub>t\<^sub>o\<^sub>r\<^sub>e = "(Location \ String.literal) list" - -fun show\<^sub>S\<^sub>t\<^sub>o\<^sub>r\<^sub>e::"'a Store \ 'a fset" where - "show\<^sub>S\<^sub>t\<^sub>o\<^sub>r\<^sub>e s = (fmran (mapping s))" - -subsubsection\Load: Memory\ - -datatype Data\<^sub>M\<^sub>e\<^sub>m\<^sub>o\<^sub>r\<^sub>y = MArray "Data\<^sub>M\<^sub>e\<^sub>m\<^sub>o\<^sub>r\<^sub>y list" - | MBool bool - | MInt int - | MAddress Address - -fun - loadRec\<^sub>M\<^sub>e\<^sub>m\<^sub>o\<^sub>r\<^sub>y :: "Location \ Data\<^sub>M\<^sub>e\<^sub>m\<^sub>o\<^sub>r\<^sub>y \ MemoryT \ MemoryT" where -"loadRec\<^sub>M\<^sub>e\<^sub>m\<^sub>o\<^sub>r\<^sub>y loc (MArray dat) mem = -(fst (foldl (\ S d . let (s',x) = S in (loadRec\<^sub>M\<^sub>e\<^sub>m\<^sub>o\<^sub>r\<^sub>y (hash loc (ShowL\<^sub>n\<^sub>a\<^sub>t x)) d s', Suc x)) (updateStore loc (MPointer loc) mem,0) dat))" -| "loadRec\<^sub>M\<^sub>e\<^sub>m\<^sub>o\<^sub>r\<^sub>y loc (MBool b) mem = updateStore loc ((MValue \ ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l) b) mem " -| "loadRec\<^sub>M\<^sub>e\<^sub>m\<^sub>o\<^sub>r\<^sub>y loc (MInt i) mem = updateStore loc ((MValue \ ShowL\<^sub>i\<^sub>n\<^sub>t) i) mem " -| "loadRec\<^sub>M\<^sub>e\<^sub>m\<^sub>o\<^sub>r\<^sub>y loc (MAddress ad) mem = updateStore loc (MValue ad) mem" - -definition load\<^sub>M\<^sub>e\<^sub>m\<^sub>o\<^sub>r\<^sub>y :: "Data\<^sub>M\<^sub>e\<^sub>m\<^sub>o\<^sub>r\<^sub>y list \ MemoryT \ MemoryT" where -"load\<^sub>M\<^sub>e\<^sub>m\<^sub>o\<^sub>r\<^sub>y dat mem = (let l = ShowL\<^sub>n\<^sub>a\<^sub>t (toploc mem); - (m, _) = foldl (\ (m',x) d . (loadRec\<^sub>M\<^sub>e\<^sub>m\<^sub>o\<^sub>r\<^sub>y (((hash l) \ ShowL\<^sub>n\<^sub>a\<^sub>t) x) d m', Suc x)) (mem, 0) dat - in (snd \ allocate) m)" - - -fun dumprec\<^sub>M\<^sub>e\<^sub>m\<^sub>o\<^sub>r\<^sub>y:: "Location \ MTypes \ MemoryT \ String.literal \ String.literal \ String.literal" where -"dumprec\<^sub>M\<^sub>e\<^sub>m\<^sub>o\<^sub>r\<^sub>y loc tp mem ls str = ( case accessStore loc mem of - Some (MPointer l) \ ( case tp of - (MTArray x t) \ iter (\ i str' . dumprec\<^sub>M\<^sub>e\<^sub>m\<^sub>o\<^sub>r\<^sub>y ((hash l o ShowL\<^sub>i\<^sub>n\<^sub>t) i) t mem - (ls + (STR ''['') + (ShowL\<^sub>i\<^sub>n\<^sub>t i) + (STR '']'')) str') str x - | _ \ FAILURE) - | Some (MValue v) \ (case tp of - MTValue t \ str + ls + (STR ''=='') + dump\<^sub>V\<^sub>a\<^sub>l\<^sub>u\<^sub>e\<^sub>t\<^sub>y\<^sub>p\<^sub>e\<^sub>s t v + (STR ''\'') - | _ \ FAILURE) - | None \ FAILURE)" - -definition dump\<^sub>M\<^sub>e\<^sub>m\<^sub>o\<^sub>r\<^sub>y :: "Location \ int \ MTypes \ MemoryT \ String.literal \String.literal \String.literal" where -"dump\<^sub>M\<^sub>e\<^sub>m\<^sub>o\<^sub>r\<^sub>y loc x t mem ls str = iter (\i. dumprec\<^sub>M\<^sub>e\<^sub>m\<^sub>o\<^sub>r\<^sub>y ((hash loc (ShowL\<^sub>i\<^sub>n\<^sub>t i))) t mem (ls + STR ''['' + (ShowL\<^sub>i\<^sub>n\<^sub>t i + STR '']''))) str x" - -subsubsection\Storage\ - -datatype Data\<^sub>S\<^sub>t\<^sub>o\<^sub>r\<^sub>a\<^sub>g\<^sub>e = - SArray "Data\<^sub>S\<^sub>t\<^sub>o\<^sub>r\<^sub>a\<^sub>g\<^sub>e list" | - SMap "(String.literal \ Data\<^sub>S\<^sub>t\<^sub>o\<^sub>r\<^sub>a\<^sub>g\<^sub>e) list" | - SBool bool | - SInt int | - SAddress Address - -definition splitAt::"nat \ String.literal \ String.literal \ String.literal" where -"splitAt n xs = (String.implode(take n (String.explode xs)), String.implode(drop n (String.explode xs)))" - - -fun splitOn':: "'a \ 'a list \ 'a list \ 'a list list" where - "splitOn' x [] acc = [rev acc]" - | "splitOn' x (y#ys) acc = (if x = y then (rev acc)#(splitOn' x ys []) - else splitOn' x ys (y#acc))" - -fun splitOn::"'a \ 'a list \ 'a list list" where -"splitOn x xs = splitOn' x xs []" - -definition isSuffixOf::"String.literal \ String.literal \ bool" where -"isSuffixOf s x = suffix (String.explode s) (String.explode x)" - -definition tolist :: "Location \ String.literal list" where -"tolist s = map String.implode (splitOn (CHR ''.'') (String.explode s))" - -abbreviation convert :: "Location \ Location" - where "convert loc \ (if loc= STR ''True'' then STR ''true'' else - if loc=STR ''False'' then STR ''false'' else loc)" - -fun go\<^sub>S\<^sub>t\<^sub>o\<^sub>r\<^sub>a\<^sub>g\<^sub>e :: "Location \ (String.literal \ STypes) \ (String.literal \ STypes)" where - "go\<^sub>S\<^sub>t\<^sub>o\<^sub>r\<^sub>a\<^sub>g\<^sub>e l (s, STArray _ t) = (s + (STR ''['') + (convert l) + (STR '']''), t)" -| "go\<^sub>S\<^sub>t\<^sub>o\<^sub>r\<^sub>a\<^sub>g\<^sub>e l (s, STMap _ t) = (s + (STR ''['') + (convert l) + (STR '']''), t)" -| "go\<^sub>S\<^sub>t\<^sub>o\<^sub>r\<^sub>a\<^sub>g\<^sub>e l (s, STValue t) = (s + (STR ''['') + (convert l) + (STR '']''), STValue t)" - -fun dumpSingle\<^sub>S\<^sub>t\<^sub>o\<^sub>r\<^sub>a\<^sub>g\<^sub>e :: "StorageT \ String.literal \ STypes \ (Location \ Location) \ String.literal \ String.literal" where -"dumpSingle\<^sub>S\<^sub>t\<^sub>o\<^sub>r\<^sub>a\<^sub>g\<^sub>e sto id' tp (loc,l) str = - (case foldr go\<^sub>S\<^sub>t\<^sub>o\<^sub>r\<^sub>a\<^sub>g\<^sub>e (tolist loc) (str + id', tp) of - (s, STValue t) \ (case fmlookup sto (loc + l) of - Some v \ s + (STR ''=='') + dump\<^sub>V\<^sub>a\<^sub>l\<^sub>u\<^sub>e\<^sub>t\<^sub>y\<^sub>p\<^sub>e\<^sub>s t v - | None \ FAILURE) - | _ \ FAILURE)" - - -definition \sorted_list_of_set' \ map_fun id id (folding_on.F insort [])\ - -lemma sorted_list_of_fset'_def': \sorted_list_of_set' = sorted_list_of_set\ - apply(rule ext) - by (simp add: sorted_list_of_set'_def sorted_list_of_set_def sorted_key_list_of_set_def) - -lemma sorted_list_of_set_sort_remdups' [code]: - \sorted_list_of_set' (set xs) = sort (remdups xs)\ - using sorted_list_of_fset'_def' sorted_list_of_set_sort_remdups - by metis - -definition locations_map :: "Location \ (Location, 'v) fmap \ Location list" where -"locations_map loc = (filter (isSuffixOf ((STR ''.'')+loc))) \ sorted_list_of_set' \ fset \ fmdom" - -definition locations :: "Location \ 'v Store \ Location list" where -"locations loc = locations_map loc \ mapping" - -fun dump\<^sub>S\<^sub>t\<^sub>o\<^sub>r\<^sub>a\<^sub>g\<^sub>e :: "StorageT \ Location \ String.literal \ STypes \ String.literal \ String.literal" -where -"dump\<^sub>S\<^sub>t\<^sub>o\<^sub>r\<^sub>a\<^sub>g\<^sub>e sto loc id' (STArray _ t) str = foldl - (\ s l . dumpSingle\<^sub>S\<^sub>t\<^sub>o\<^sub>r\<^sub>a\<^sub>g\<^sub>e sto id' t ((splitAt (length (String.explode l) - length (String.explode loc) - 1) l)) s - + (STR ''\'')) str (locations_map loc sto)" -| "dump\<^sub>S\<^sub>t\<^sub>o\<^sub>r\<^sub>a\<^sub>g\<^sub>e sto loc id' (STMap _ t) str = - foldl (\ s l . dumpSingle\<^sub>S\<^sub>t\<^sub>o\<^sub>r\<^sub>a\<^sub>g\<^sub>e sto id' t (splitAt (length (String.explode l) - length (String.explode loc) - 1) l) s + (STR ''\'')) str -(locations_map loc sto)" - | "dump\<^sub>S\<^sub>t\<^sub>o\<^sub>r\<^sub>a\<^sub>g\<^sub>e sto loc id' (STValue t) str = (case fmlookup sto loc of - Some v \ str + id' + (STR ''=='') + dump\<^sub>V\<^sub>a\<^sub>l\<^sub>u\<^sub>e\<^sub>t\<^sub>y\<^sub>p\<^sub>e\<^sub>s t v + (STR ''\'') - | _ \ str)" - - -fun loadRec\<^sub>S\<^sub>t\<^sub>o\<^sub>r\<^sub>a\<^sub>g\<^sub>e :: "Location \ Data\<^sub>S\<^sub>t\<^sub>o\<^sub>r\<^sub>a\<^sub>g\<^sub>e \ StorageT \ StorageT" where -"loadRec\<^sub>S\<^sub>t\<^sub>o\<^sub>r\<^sub>a\<^sub>g\<^sub>e loc (SArray dat) sto = fst (foldl (\ S d . let (s', x) = S in (loadRec\<^sub>S\<^sub>t\<^sub>o\<^sub>r\<^sub>a\<^sub>g\<^sub>e (hash loc (ShowL\<^sub>n\<^sub>a\<^sub>t x)) d s', Suc x)) (sto,0) dat)" -| "loadRec\<^sub>S\<^sub>t\<^sub>o\<^sub>r\<^sub>a\<^sub>g\<^sub>e loc (SMap dat) sto = ( foldr (\ (k, v) s'. loadRec\<^sub>S\<^sub>t\<^sub>o\<^sub>r\<^sub>a\<^sub>g\<^sub>e (hash loc k) v s') dat sto)" -| "loadRec\<^sub>S\<^sub>t\<^sub>o\<^sub>r\<^sub>a\<^sub>g\<^sub>e loc (SBool b) sto = fmupd loc (ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l b) sto" -| "loadRec\<^sub>S\<^sub>t\<^sub>o\<^sub>r\<^sub>a\<^sub>g\<^sub>e loc (SInt i) sto = fmupd loc (ShowL\<^sub>i\<^sub>n\<^sub>t i) sto" -| "loadRec\<^sub>S\<^sub>t\<^sub>o\<^sub>r\<^sub>a\<^sub>g\<^sub>e loc (SAddress ad) sto = fmupd loc ad sto" - -subsubsection\Environment\ - -datatype Data\<^sub>E\<^sub>n\<^sub>v\<^sub>i\<^sub>r\<^sub>o\<^sub>n\<^sub>m\<^sub>e\<^sub>n\<^sub>t = - Memarr "Data\<^sub>M\<^sub>e\<^sub>m\<^sub>o\<^sub>r\<^sub>y list" | - CDarr "Data\<^sub>M\<^sub>e\<^sub>m\<^sub>o\<^sub>r\<^sub>y list" | - Stoarr "Data\<^sub>S\<^sub>t\<^sub>o\<^sub>r\<^sub>a\<^sub>g\<^sub>e list"| - Stomap "(String.literal \ Data\<^sub>S\<^sub>t\<^sub>o\<^sub>r\<^sub>a\<^sub>g\<^sub>e) list" | - Stackbool bool | - Stobool bool | - Stackint int | - Stoint int | - Stackaddr Address | - Stoaddr Address - -fun loadsimple\<^sub>E\<^sub>n\<^sub>v\<^sub>i\<^sub>r\<^sub>o\<^sub>n\<^sub>m\<^sub>e\<^sub>n\<^sub>t :: "(Stack \ CalldataT \ MemoryT \ StorageT \ Environment) - \ (Identifier \ Type \ Data\<^sub>E\<^sub>n\<^sub>v\<^sub>i\<^sub>r\<^sub>o\<^sub>n\<^sub>m\<^sub>e\<^sub>n\<^sub>t) \ (Stack \ CalldataT \ MemoryT \ StorageT \ Environment)" - where -"loadsimple\<^sub>E\<^sub>n\<^sub>v\<^sub>i\<^sub>r\<^sub>o\<^sub>n\<^sub>m\<^sub>e\<^sub>n\<^sub>t (k, c, m, s, e) (id', tp, d) = (case d of - Stackbool b \ - let (k', e') = astack id' tp (KValue (ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l b)) (k, e) - in (k', c, m, s, e') - | Stobool b \ - let (s', e') = astore id' tp (ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l b) (s, e) - in (k, c, m, s', e') - | Stackint n \ - let (k', e') = astack id' tp (KValue (ShowL\<^sub>i\<^sub>n\<^sub>t n)) (k, e) - in (k', c, m, s, e') - | Stoint n \ - let (s', e') = astore id' tp (ShowL\<^sub>i\<^sub>n\<^sub>t n) (s, e) - in (k, c, m, s', e') - | Stackaddr ad \ - let (k', e') = astack id' tp (KValue ad) (k, e) - in (k', c, m, s, e') - | Stoaddr ad \ - let (s', e') = astore id' tp ad (s, e) - in (k, c, m, s', e') - | CDarr a \ - let l = ShowL\<^sub>n\<^sub>a\<^sub>t (toploc c); - c' = load\<^sub>M\<^sub>e\<^sub>m\<^sub>o\<^sub>r\<^sub>y a c; - (k', e') = astack id' tp (KCDptr l) (k, e) - in (k', c', m, s, e') - | Memarr a \ - let l = ShowL\<^sub>n\<^sub>a\<^sub>t (toploc m); - m' = load\<^sub>M\<^sub>e\<^sub>m\<^sub>o\<^sub>r\<^sub>y a m; - (k', e') = astack id' tp (KMemptr l) (k, e) - in (k', c, m', s, e') - | Stoarr a \ - let s' = loadRec\<^sub>S\<^sub>t\<^sub>o\<^sub>r\<^sub>a\<^sub>g\<^sub>e id' (SArray a) s; - e' = updateEnv id' tp (Storeloc id') e - in (k, c, m, s', e') - | Stomap mp \ - let s' = loadRec\<^sub>S\<^sub>t\<^sub>o\<^sub>r\<^sub>a\<^sub>g\<^sub>e id' (SMap mp) s; - e' = updateEnv id' tp (Storeloc id') e - in (k, c, m, s', e') -)" - -definition load\<^sub>E\<^sub>n\<^sub>v\<^sub>i\<^sub>r\<^sub>o\<^sub>n\<^sub>m\<^sub>e\<^sub>n\<^sub>t::"(Stack \ CalldataT \ MemoryT \ StorageT \ Environment) \ (Identifier \ Type \ Data\<^sub>E\<^sub>n\<^sub>v\<^sub>i\<^sub>r\<^sub>o\<^sub>n\<^sub>m\<^sub>e\<^sub>n\<^sub>t) list - \ (Stack \ CalldataT \ MemoryT \ StorageT \ Environment)" - where -"load\<^sub>E\<^sub>n\<^sub>v\<^sub>i\<^sub>r\<^sub>o\<^sub>n\<^sub>m\<^sub>e\<^sub>n\<^sub>t = foldl loadsimple\<^sub>E\<^sub>n\<^sub>v\<^sub>i\<^sub>r\<^sub>o\<^sub>n\<^sub>m\<^sub>e\<^sub>n\<^sub>t" - -definition getValue\<^sub>E\<^sub>n\<^sub>v\<^sub>i\<^sub>r\<^sub>o\<^sub>n\<^sub>m\<^sub>e\<^sub>n\<^sub>t :: "Stack \ CalldataT \ MemoryT \ StorageT \ Environment \ Identifier \ String.literal \ String.literal" - where -"getValue\<^sub>E\<^sub>n\<^sub>v\<^sub>i\<^sub>r\<^sub>o\<^sub>n\<^sub>m\<^sub>e\<^sub>n\<^sub>t k c m s e i txt = (case fmlookup (denvalue e) i of - Some (tp, Stackloc l) \ (case accessStore l k of - Some (KValue v) \ (case tp of - Value t \ (txt + i + (STR ''=='') + dump\<^sub>V\<^sub>a\<^sub>l\<^sub>u\<^sub>e\<^sub>t\<^sub>y\<^sub>p\<^sub>e\<^sub>s t v + (STR ''\'')) - | _ \ FAILURE) - | Some (KCDptr p) \ (case tp of - Calldata (MTArray x t) \ dump\<^sub>M\<^sub>e\<^sub>m\<^sub>o\<^sub>r\<^sub>y p x t c i txt - | _ \ FAILURE) - | Some (KMemptr p) \ (case tp of - Memory (MTArray x t) \ dump\<^sub>M\<^sub>e\<^sub>m\<^sub>o\<^sub>r\<^sub>y p x t m i txt - | _ \ FAILURE) - | Some (KStoptr p) \ (case tp of - Storage t \ dump\<^sub>S\<^sub>t\<^sub>o\<^sub>r\<^sub>a\<^sub>g\<^sub>e s p i t txt - | _ \ FAILURE)) - | Some (Storage t, Storeloc l) \ dump\<^sub>S\<^sub>t\<^sub>o\<^sub>r\<^sub>a\<^sub>g\<^sub>e s l i t txt - | _ \ FAILURE -)" - -type_synonym Data\<^sub>P = "(Address \ ((Identifier \ Member) list \ S))" - -definition dump\<^sub>E\<^sub>n\<^sub>v\<^sub>i\<^sub>r\<^sub>o\<^sub>n\<^sub>m\<^sub>e\<^sub>n\<^sub>t :: "Stack \ CalldataT \ MemoryT \ StorageT \ Environment \ Identifier list \ String.literal" - where "dump\<^sub>E\<^sub>n\<^sub>v\<^sub>i\<^sub>r\<^sub>o\<^sub>n\<^sub>m\<^sub>e\<^sub>n\<^sub>t k c m s e sl = foldr (getValue\<^sub>E\<^sub>n\<^sub>v\<^sub>i\<^sub>r\<^sub>o\<^sub>n\<^sub>m\<^sub>e\<^sub>n\<^sub>t k c m s e) sl (STR '''')" - -fun loadProc::"Environment\<^sub>P \ Data\<^sub>P \ Environment\<^sub>P" - where "loadProc e\<^sub>P (ad, (xs, fb)) = fmupd ad (fmap_of_list xs, fb) e\<^sub>P" - -fun initStorage::"(Address \ Balance) list \ (Address, StorageT) fmap \ (Address, StorageT) fmap" - where "initStorage [] m = m" - | "initStorage (x # xs) m = fmupd (fst x) (fmempty) m" - -subsection\Test Setup\ - -definition eval::"Gas \ (S \ Environment\<^sub>P \ Environment \ CalldataT \ (unit, Ex, State) state_monad) - \ S \ Address \ Address \ Valuetype \ (Address \ Balance) list - \ Data\<^sub>P list - \ (String.literal \ Type \ Data\<^sub>E\<^sub>n\<^sub>v\<^sub>i\<^sub>r\<^sub>o\<^sub>n\<^sub>m\<^sub>e\<^sub>n\<^sub>t) list - \ String.literal" - where "eval g stmt\<^sub>e\<^sub>v\<^sub>a\<^sub>l stm addr adest aval acc d dat - = (let (k,c,m,s,e) = load\<^sub>E\<^sub>n\<^sub>v\<^sub>i\<^sub>r\<^sub>o\<^sub>n\<^sub>m\<^sub>e\<^sub>n\<^sub>t (emptyStore, emptyStore, emptyStore, fmempty, emptyEnv addr adest aval) dat; - e\<^sub>p = foldl loadProc fmempty d; - a = init\<^sub>A\<^sub>c\<^sub>c\<^sub>o\<^sub>u\<^sub>n\<^sub>t acc; - s' = fmupd addr s (initStorage acc fmempty); - z = \accounts=a,stack=k,memory=m,storage=s',gas=g\ - in ( - case (stmt\<^sub>e\<^sub>v\<^sub>a\<^sub>l stm e\<^sub>p e c z) of - Normal ((), z') \ (dump\<^sub>E\<^sub>n\<^sub>v\<^sub>i\<^sub>r\<^sub>o\<^sub>n\<^sub>m\<^sub>e\<^sub>n\<^sub>t (stack z') c (memory z') (the (fmlookup (storage z') addr)) e (map (\ (a,b,c). a) dat)) - + (dump\<^sub>A\<^sub>c\<^sub>c\<^sub>o\<^sub>u\<^sub>n\<^sub>t\<^sub>s (accounts z') (map fst acc)) - | Exception Err \ STR ''Exception'' - | Exception Gas \ STR ''OutOfGas''))" - -value "eval 1 - stmt - SKIP - (STR ''089Be5381FcEa58aF334101414c04F993947C733'') - (STR '''') - (STR ''0'') - [(STR ''089Be5381FcEa58aF334101414c04F993947C733'', STR ''100''),(STR ''115f6e2F70210C14f7DB1AC69737a3CC78435d49'', STR ''100'')] - [] - [(STR ''v1'', (Value TBool, Stackbool True))]" - -lemma "eval 1000 - stmt - SKIP - (STR ''089Be5381FcEa58aF334101414c04F993947C733'') - (STR '''') - (STR ''0'') - [(STR ''089Be5381FcEa58aF334101414c04F993947C733'', STR ''100''), (STR ''115f6e2F70210C14f7DB1AC69737a3CC78435d49'', STR ''100'')] - [] - [(STR ''v1'', (Value TBool, Stackbool True))] - = STR ''v1==true\089Be5381FcEa58aF334101414c04F993947C733.balance==100\115f6e2F70210C14f7DB1AC69737a3CC78435d49.balance==100\''" - by(eval) - -value "eval 1000 - stmt - SKIP - (STR ''089Be5381FcEa58aF334101414c04F993947C733'') - (STR '''') - (STR ''0'') - [(STR ''089Be5381FcEa58aF334101414c04F993947C733'', STR ''100''), (STR ''115f6e2F70210C14f7DB1AC69737a3CC78435d49'', STR ''100'')] - [] - [(STR ''v1'', (Storage (STArray 5 (STValue TBool)), Stoarr [SBool True, SBool False, SBool True, SBool False, SBool True]))]" - - -lemma "eval 1000 - stmt - SKIP - (STR ''089Be5381FcEa58aF334101414c04F993947C733'') - (STR '''') - (STR ''0'') - [(STR ''089Be5381FcEa58aF334101414c04F993947C733'', STR ''100''), (STR ''115f6e2F70210C14f7DB1AC69737a3CC78435d49'', STR ''100'')] - [] - [(STR ''v1'',(Memory (MTArray 5 (MTValue TBool)), Memarr [MBool True, MBool False, MBool True, MBool False, MBool True]))] - = STR ''v1[0]==true\v1[1]==false\v1[2]==true\v1[3]==false\v1[4]==true\089Be5381FcEa58aF334101414c04F993947C733.balance==100\115f6e2F70210C14f7DB1AC69737a3CC78435d49.balance==100\''" - by(eval) - -lemma "eval 1000 - stmt - (ITE FALSE (ASSIGN (Id (STR ''x'')) TRUE) (ASSIGN (Id (STR ''y'')) TRUE)) - (STR ''089Be5381FcEa58aF334101414c04F993947C733'') - (STR '''') - (STR ''0'') - [(STR ''089Be5381FcEa58aF334101414c04F993947C733'', STR ''100''), (STR ''115f6e2F70210C14f7DB1AC69737a3CC78435d49'', STR ''100'')] - [] - [(STR ''x'', (Value TBool, Stackbool False)), (STR ''y'', (Value TBool, Stackbool False))] - = STR ''y==true\x==false\089Be5381FcEa58aF334101414c04F993947C733.balance==100\115f6e2F70210C14f7DB1AC69737a3CC78435d49.balance==100\''" - by(eval) - -lemma "eval 1000 - stmt - (BLOCK ((STR ''v2'', Value TBool), None) (ASSIGN (Id (STR ''v1'')) (LVAL (Id (STR ''v2''))))) - (STR ''089Be5381FcEa58aF334101414c04F993947C733'') - (STR '''') - (STR ''0'') - [(STR ''089Be5381FcEa58aF334101414c04F993947C733'', STR ''100''), (STR ''115f6e2F70210C14f7DB1AC69737a3CC78435d49'', STR ''100'')] - [] - [(STR ''v1'', (Value TBool, Stackbool True))] - = STR ''v1==false\089Be5381FcEa58aF334101414c04F993947C733.balance==100\115f6e2F70210C14f7DB1AC69737a3CC78435d49.balance==100\''" - by(eval) - -lemma "eval 1000 - stmt - (ASSIGN (Id (STR ''a_s120_21_m8'')) (LVAL (Id (STR ''a_s120_21_s8'')))) - (STR ''089Be5381FcEa58aF334101414c04F993947C733'') - (STR '''') - (STR ''0'') - [(STR ''089Be5381FcEa58aF334101414c04F993947C733'', STR ''100'')] - [] - [((STR ''a_s120_21_s8''), Storage (STArray 1 (STArray 2 (STValue (TSInt 120)))), Stoarr [SArray [SInt 347104507864064359095275590289383142, SInt 565831699297331399489670920129618233]]), - ((STR ''a_s120_21_m8''), Memory (MTArray 1 (MTArray 2 (MTValue (TSInt 120)))), Memarr [MArray [MInt (290845675805142398428016622247257774), MInt ((-96834026877269277170645294669272226))]])] -= STR ''a_s120_21_m8[0][0]==347104507864064359095275590289383142\a_s120_21_m8[0][1]==565831699297331399489670920129618233\a_s120_21_s8[0][0]==347104507864064359095275590289383142\a_s120_21_s8[0][1]==565831699297331399489670920129618233\089Be5381FcEa58aF334101414c04F993947C733.balance==100\''" - by(eval) - -lemma "eval 1000 - stmt - (ASSIGN (Ref (STR ''a_s8_32_m0'') [UINT 8 1]) (LVAL (Ref (STR ''a_s8_31_s7'') [UINT 8 0]))) - (STR ''089Be5381FcEa58aF334101414c04F993947C733'') - (STR '''') - (STR ''0'') - [(STR ''089Be5381FcEa58aF334101414c04F993947C733'', STR ''100'')] - [] - [(STR ''a_s8_31_s7'', (Storage (STArray 1 (STArray 3 (STValue (TSInt 8)))), Stoarr [SArray [SInt ((98)), SInt ((-23)), SInt (36)]])), - (STR ''a_s8_32_m0'', (Memory (MTArray 2 (MTArray 3 (MTValue (TSInt 8)))), Memarr [MArray [MInt ((-64)), MInt ((39)), MInt ((-125))], MArray [MInt ((-32)), MInt ((-82)), MInt ((-105))]]))] - = STR ''a_s8_32_m0[0][0]==-64\a_s8_32_m0[0][1]==39\a_s8_32_m0[0][2]==-125\a_s8_32_m0[1][0]==98\a_s8_32_m0[1][1]==-23\a_s8_32_m0[1][2]==36\a_s8_31_s7[0][0]==98\a_s8_31_s7[0][1]==-23\a_s8_31_s7[0][2]==36\089Be5381FcEa58aF334101414c04F993947C733.balance==100\''" - by(eval) - - -lemma "eval 1000 - stmt - SKIP - (STR ''089Be5381FcEa58aF334101414c04F993947C733'') - (STR '''') - (STR ''0'') - [(STR ''089Be5381FcEa58aF334101414c04F993947C733'', STR ''100''), (STR ''115f6e2F70210C14f7DB1AC69737a3CC78435d49'', STR ''100'')] - [] - [(STR ''v1'', (Storage (STMap (TUInt 32) (STValue (TUInt 8))), Stomap [(STR ''2129136830'', SInt (247))]))] - = STR ''v1[2129136830]==247\089Be5381FcEa58aF334101414c04F993947C733.balance==100\115f6e2F70210C14f7DB1AC69737a3CC78435d49.balance==100\''" - by(eval) - -value "eval 1000 - stmt - (INVOKE (STR ''m1'') []) - (STR ''myaddr'') - (STR '''') - (STR ''0'') - [(STR ''myaddr'', STR ''100'')] - [ - (STR ''myaddr'', - ([(STR ''m1'', Method ([], SKIP, None))], - SKIP)) - ] - [(STR ''x'', (Value TBool, Stackbool True))]" - -lemma "eval 1000 - stmt - (ASSIGN (Id (STR ''v1'')) (CALL (STR ''m1'') [])) - (STR ''myaddr'') - (STR '''') - (STR ''0'') - [(STR ''myaddr'', STR ''100'')] - [ - (STR ''myaddr'', - ([(STR ''m1'', Method ([], SKIP, Some (UINT 8 5)))], - SKIP)) - ] - [(STR ''v1'', (Value (TUInt 8), Stackint 0))] - = STR ''v1==5\myaddr.balance==100\''" - by (eval) - -lemma "eval 1000 - stmt - (ASSIGN (Id (STR ''v1'')) (CALL (STR ''m1'') [E.INT 8 3, E.INT 8 4])) - (STR ''myaddr'') - (STR '''') - (STR ''0'') - [(STR ''myaddr'', STR ''100'')] - [ - (STR ''myaddr'', - ([(STR ''m1'', Method ([(STR ''v2'', Value (TSInt 8)), (STR ''v3'', Value (TSInt 8))], SKIP, Some (PLUS (LVAL (Id (STR ''v2''))) (LVAL (Id (STR ''v3''))))))], - SKIP)) - ] - [(STR ''v1'', (Value (TSInt 8), Stackint 0))] - = STR ''v1==7\myaddr.balance==100\''" - by (eval) - -lemma "eval 1000 - stmt - (ASSIGN (Id (STR ''v1'')) (ECALL (ADDRESS (STR ''extaddr'')) (STR ''m1'') [E.INT 8 3, E.INT 8 4] (E.UINT 8 0))) - (STR ''myaddr'') - (STR '''') - (STR ''0'') - [(STR ''myaddr'', STR ''100'')] - [ - (STR ''extaddr'', - ([(STR ''m1'', Method ([(STR ''v2'', Value (TSInt 8)), (STR ''v3'', Value (TSInt 8))], SKIP, Some (PLUS (LVAL (Id (STR ''v2''))) (LVAL (Id (STR ''v3''))))))], - SKIP)) - ] - [(STR ''v1'', (Value (TSInt 8), Stackint 0))] - = STR ''v1==7\myaddr.balance==100\''" - by (eval) - -lemma "eval 1000 - stmt - (TRANSFER (ADDRESS (STR ''myaddr'')) (UINT 256 10)) - (STR ''089Be5381FcEa58aF334101414c04F993947C733'') - (STR '''') - (STR ''0'') - [(STR ''089Be5381FcEa58aF334101414c04F993947C733'', STR ''100''), (STR ''0x2d5F6f401c770eEAdd68deB348948ed4504c4676'', STR ''100'')] - [ - (STR ''myaddr'', - ([], SKIP)) - ] - [] - = STR ''089Be5381FcEa58aF334101414c04F993947C733.balance==90\0x2d5F6f401c770eEAdd68deB348948ed4504c4676.balance==100\''" - by (eval) - -value "eval 1000 - stmt - (TRANSFER (ADDRESS (STR ''myaddr'')) (UINT 256 10)) - (STR ''089Be5381FcEa58aF334101414c04F993947C733'') - (STR '''') - (STR ''0'') - [(STR ''089Be5381FcEa58aF334101414c04F993947C733'', STR ''100''), (STR ''0x2d5F6f401c770eEAdd68deB348948ed4504c4676'', STR ''100'')] - [ - (STR ''myaddr'', - ([], SKIP)) - ] - []" - -lemma "eval 1000 - stmt - (COMP(COMP(((ASSIGN (Id (STR ''x'')) (E.UINT 8 0))))(TRANSFER (ADDRESS (STR ''myaddr'')) (UINT 256 5)))(SKIP)) - (STR ''089Be5381FcEa58aF334101414c04F993947C733'') - (STR '''') - (STR ''0'') - [(STR ''089Be5381FcEa58aF334101414c04F993947C733'', STR ''100''), (STR ''115f6e2F70210C14f7DB1AC69737a3CC78435d49'', STR ''100'')] - [ - (STR ''myaddr'', - ([], SKIP)) - ] - [(STR ''x'', (Value (TUInt 8), Stackint 9))] - = STR ''x==0\089Be5381FcEa58aF334101414c04F993947C733.balance==95\115f6e2F70210C14f7DB1AC69737a3CC78435d49.balance==100\''" - by (eval) - -value "eval 1000 - stmt - (EXTERNAL (ADDRESS (STR ''Victim'')) (STR ''withdraw'') [] (E.UINT 8 0)) - (STR ''Victim'') - (STR '''') - (STR ''0'') - [(STR ''Victim'', STR ''100''), (STR ''Attacker'', STR ''100'')] - [ - (STR ''Attacker'', - [(STR ''withdraw'', Method ([], EXTERNAL (ADDRESS (STR ''Victim'')) (STR ''withdraw'') [] (E.UINT 8 0), None))], - SKIP), - (STR ''Victim'', - [(STR ''withdraw'', Method ([], EXTERNAL (ADDRESS (STR ''Attacker'')) (STR ''withdraw'') [] (E.UINT 8 0), None))], - SKIP) - ] - []" - -value "eval 1000 - stmt - (INVOKE (STR ''withdraw'') []) - (STR ''Victim'') - (STR '''') - (STR ''0'') - [(STR ''Victim'', STR ''100''), (STR ''Attacker'', STR ''100'')] - [ - (STR ''Victim'', - [(STR ''withdraw'', Method ([], INVOKE (STR ''withdraw'') [], None))], - SKIP) - ] - []" - -subsection\The Final Code Export\ - -consts ReadL\<^sub>S :: "String.literal \ S" -consts ReadL\<^sub>a\<^sub>c\<^sub>c :: "String.literal \ (String.literal \ String.literal) list" -consts ReadL\<^sub>d\<^sub>a\<^sub>t :: "String.literal \ (String.literal \ Type \ Data\<^sub>E\<^sub>n\<^sub>v\<^sub>i\<^sub>r\<^sub>o\<^sub>n\<^sub>m\<^sub>e\<^sub>n\<^sub>t) list" -consts ReadL\<^sub>P :: "String.literal \ Data\<^sub>P list" - -code_printing - constant ReadL\<^sub>S \ (Haskell) "Prelude.read" - | constant ReadL\<^sub>a\<^sub>c\<^sub>c \ (Haskell) "Prelude.read" - | constant ReadL\<^sub>d\<^sub>a\<^sub>t \ (Haskell) "Prelude.read" - | constant ReadL\<^sub>P \ (Haskell) "Prelude.read" - -fun main_stub :: "String.literal list \ (int \ String.literal)" - where - "main_stub [credit, stm, saddr, raddr, val, acc, pr, dat] - = (0, eval (ReadL\<^sub>n\<^sub>a\<^sub>t credit) stmt (ReadL\<^sub>S stm) saddr raddr val (ReadL\<^sub>a\<^sub>c\<^sub>c acc) (ReadL\<^sub>P pr) (ReadL\<^sub>d\<^sub>a\<^sub>t dat))" - | "main_stub [stm, saddr, raddr, val, acc, pr, dat] - = (0, eval 1000 stmt (ReadL\<^sub>S stm) saddr raddr val (ReadL\<^sub>a\<^sub>c\<^sub>c acc) (ReadL\<^sub>P pr) (ReadL\<^sub>d\<^sub>a\<^sub>t dat))" - | "main_stub _ = (2, - STR ''solidity-evaluator [credit] \"Statement\" \"ContractAddress\" \"OriginAddress\" \"Value\"\'' - + STR '' \"(Address * Balance) list\" \"(Address * ((Identifier * Member) list) * Statement)\" \"(Variable * Type * Value) list\"\'' - + STR ''\'')" - -generate_file "code/solidity-evaluator/app/Main.hs" = \ -module Main where -import System.Environment -import Solidity_Evaluator -import Prelude - -main :: IO () -main = do - args <- getArgs - Prelude.putStr(snd $ Solidity_Evaluator.main_stub args) -\ - -export_generated_files _ - -export_code eval SKIP main_stub - in Haskell module_name "Solidity_Evaluator" file_prefix "solidity-evaluator/src" (string_classes) - -subsection\Demonstrating the Symbolic Execution of Solidity\ - -abbreviation P1::S - where "P1 \ COMP (ASSIGN (Id (STR ''sa'')) (LVAL (Id (STR ''ma'')))) - (ASSIGN (Ref (STR ''sa'') [UINT (8::nat) 0]) TRUE)" -abbreviation myenv::Environment - where "myenv \ updateEnv (STR ''ma'') (Memory (MTArray 1 (MTValue TBool))) (Stackloc (STR ''1'')) - (updateEnv (STR ''sa'') (Storage (STArray 1 (STValue TBool))) (Storeloc (STR ''1'')) - (emptyEnv (STR ''ad'') (STR ''ad'') (STR ''0'')))" - -abbreviation mystack::Stack - where "mystack \ updateStore (STR ''1'') (KMemptr (STR ''1'')) emptyStore" - -abbreviation mystore::StorageT - where "mystore \ fmempty" - -abbreviation mymemory::MemoryT - where "mymemory \ updateStore (STR ''0.1'') (MValue (STR ''false'')) emptyStore" - -abbreviation mystorage::StorageT - where "mystorage \ fmupd (STR ''0.1'') (STR ''True'') fmempty" - - -lemma "( stmt P1 fmempty myenv emptyStore \accounts=emptyAccount, stack=mystack, memory=mymemory, storage=fmupd (STR ''ad'') mystorage fmempty, gas=1000\ ) = - (Normal ((), \accounts=emptyAccount, stack=mystack, memory=mymemory, storage=fmupd (STR ''ad'') mystorage fmempty, gas=1000\ ))" - by(solidity_symbex) - -end +section\ Solidty Evaluator and Code Generator Setup\ + +theory + Solidity_Evaluator +imports + Solidity_Main + "HOL-Library.Code_Target_Numeral" + "HOL-Library.Sublist" + "HOL-Library.Finite_Map" +begin + +paragraph\Generalized Unit Tests\ + +lemma "createSInt 8 500 = STR ''-12''" + by(eval) + +lemma "STR ''-92134039538802366542421159375273829975'' + = createSInt 128 45648483135649456465465452123894894554654654654654646999465" + by(eval) + +lemma "STR ''-128'' = createSInt 8 (-128)" + by(eval) + +lemma "STR ''244'' = (createUInt 8 500)" + by(eval) + +lemma "STR ''220443428915524155977936330922349307608'' + = (createUInt 128 4564848313564945646546545212389489455465465465465464699946544654654654654168)" + by(eval) + +lemma "less (TUInt 144) (TSInt 160) (STR ''5'') (STR ''8'') = Some(STR ''True'', TBool) " + by(eval) + +subsection\Code Generator Setup and Local Tests\ + +subsubsection\Utils\ + +definition EMPTY::"String.literal" where "EMPTY = STR ''''" + +definition FAILURE::"String.literal" where "FAILURE = STR ''Failure''" + +fun intersperse :: "String.literal \ String.literal list \ String.literal" where + "intersperse s [] = EMPTY" +| "intersperse s [x] = x" +| "intersperse s (x # xs) = x + s + intersperse s xs" + +definition splitAt::"nat \ String.literal \ String.literal \ String.literal" where +"splitAt n xs = (String.implode(take n (String.explode xs)), String.implode(drop n (String.explode xs)))" + +fun splitOn':: "'a \ 'a list \ 'a list \ 'a list list" where + "splitOn' x [] acc = [rev acc]" + | "splitOn' x (y#ys) acc = (if x = y then (rev acc)#(splitOn' x ys []) + else splitOn' x ys (y#acc))" + +fun splitOn::"'a \ 'a list \ 'a list list" where +"splitOn x xs = splitOn' x xs []" + +definition isSuffixOf::"String.literal \ String.literal \ bool" where +"isSuffixOf s x = suffix (String.explode s) (String.explode x)" + +definition tolist :: "Location \ String.literal list" where +"tolist s = map String.implode (splitOn (CHR ''.'') (String.explode s))" + +abbreviation convert :: "Location \ Location" + where "convert loc \ (if loc= STR ''True'' then STR ''true'' else + if loc=STR ''False'' then STR ''false'' else loc)" + +definition \sorted_list_of_set' \ map_fun id id (folding_on.F insort [])\ + +lemma sorted_list_of_fset'_def': \sorted_list_of_set' = sorted_list_of_set\ + apply(rule ext) + by (simp add: sorted_list_of_set'_def sorted_list_of_set_def sorted_key_list_of_set_def) + +lemma sorted_list_of_set_sort_remdups' [code]: + \sorted_list_of_set' (set xs) = sort (remdups xs)\ + using sorted_list_of_fset'_def' sorted_list_of_set_sort_remdups + by metis + +definition locations_map :: "Location \ (Location, 'v) fmap \ Location list" where +"locations_map loc = (filter (isSuffixOf ((STR ''.'')+loc))) \ sorted_list_of_set' \ fset \ fmdom" + +definition locations :: "Location \ 'v Store \ Location list" where +"locations loc = locations_map loc \ mapping" + +subsubsection\Valuetypes\ + +fun dump\<^sub>V\<^sub>a\<^sub>l\<^sub>u\<^sub>e\<^sub>t\<^sub>y\<^sub>p\<^sub>e\<^sub>s::"Types \ Valuetype \ String.literal" where + "dump\<^sub>V\<^sub>a\<^sub>l\<^sub>u\<^sub>e\<^sub>t\<^sub>y\<^sub>p\<^sub>e\<^sub>s (TSInt _) n = n" + | "dump\<^sub>V\<^sub>a\<^sub>l\<^sub>u\<^sub>e\<^sub>t\<^sub>y\<^sub>p\<^sub>e\<^sub>s (TUInt _) n = n" + | "dump\<^sub>V\<^sub>a\<^sub>l\<^sub>u\<^sub>e\<^sub>t\<^sub>y\<^sub>p\<^sub>e\<^sub>s TBool b = (if b = (STR ''True'') then STR ''true'' else STR ''false'')" + | "dump\<^sub>V\<^sub>a\<^sub>l\<^sub>u\<^sub>e\<^sub>t\<^sub>y\<^sub>p\<^sub>e\<^sub>s TAddr ad = ad" + +subsubsection\Memory\ + +datatype Data\<^sub>M\<^sub>e\<^sub>m\<^sub>o\<^sub>r\<^sub>y = MArray "Data\<^sub>M\<^sub>e\<^sub>m\<^sub>o\<^sub>r\<^sub>y list" + | MBool bool + | MInt int + | MAddress Address + +fun loadRec\<^sub>M\<^sub>e\<^sub>m\<^sub>o\<^sub>r\<^sub>y :: "Location \ Data\<^sub>M\<^sub>e\<^sub>m\<^sub>o\<^sub>r\<^sub>y \ MemoryT \ MemoryT" and + iterateM :: "Location \ MemoryT \ nat \ Data\<^sub>M\<^sub>e\<^sub>m\<^sub>o\<^sub>r\<^sub>y \ MemoryT \ nat" where + "loadRec\<^sub>M\<^sub>e\<^sub>m\<^sub>o\<^sub>r\<^sub>y loc (MArray dat) mem = fst (foldl (iterateM loc) (updateStore loc (MPointer loc) mem,0) dat)" +| "loadRec\<^sub>M\<^sub>e\<^sub>m\<^sub>o\<^sub>r\<^sub>y loc (MBool b) mem = updateStore loc ((MValue \ ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l) b) mem " +| "loadRec\<^sub>M\<^sub>e\<^sub>m\<^sub>o\<^sub>r\<^sub>y loc (MInt i) mem = updateStore loc ((MValue \ ShowL\<^sub>i\<^sub>n\<^sub>t) i) mem " +| "loadRec\<^sub>M\<^sub>e\<^sub>m\<^sub>o\<^sub>r\<^sub>y loc (MAddress ad) mem = updateStore loc (MValue ad) mem" +| "iterateM loc (mem,x) d = (loadRec\<^sub>M\<^sub>e\<^sub>m\<^sub>o\<^sub>r\<^sub>y (hash loc (ShowL\<^sub>n\<^sub>a\<^sub>t x)) d mem, Suc x)" + +definition load\<^sub>M\<^sub>e\<^sub>m\<^sub>o\<^sub>r\<^sub>y :: "Data\<^sub>M\<^sub>e\<^sub>m\<^sub>o\<^sub>r\<^sub>y list \ MemoryT \ MemoryT" where +"load\<^sub>M\<^sub>e\<^sub>m\<^sub>o\<^sub>r\<^sub>y dat mem = (let loc = ShowL\<^sub>n\<^sub>a\<^sub>t (toploc mem); + (m, _) = foldl (iterateM loc) (mem, 0) dat + in (snd \ allocate) m)" + +fun dumprec\<^sub>M\<^sub>e\<^sub>m\<^sub>o\<^sub>r\<^sub>y:: "Location \ MTypes \ MemoryT \ String.literal \ String.literal \ String.literal" where +"dumprec\<^sub>M\<^sub>e\<^sub>m\<^sub>o\<^sub>r\<^sub>y loc tp mem ls str = + (case accessStore loc mem of + Some (MPointer l) \ + (case tp of + (MTArray x t) \ iter (\i str' . dumprec\<^sub>M\<^sub>e\<^sub>m\<^sub>o\<^sub>r\<^sub>y ((hash l o ShowL\<^sub>i\<^sub>n\<^sub>t) i) t mem + (ls + (STR ''['') + (ShowL\<^sub>i\<^sub>n\<^sub>t i) + (STR '']'')) str') str x + | _ \ FAILURE) + | Some (MValue v) \ + (case tp of + MTValue t \ str + ls + (STR ''=='') + dump\<^sub>V\<^sub>a\<^sub>l\<^sub>u\<^sub>e\<^sub>t\<^sub>y\<^sub>p\<^sub>e\<^sub>s t v + (STR ''\'') + | _ \ FAILURE) + | None \ FAILURE)" + +definition dump\<^sub>M\<^sub>e\<^sub>m\<^sub>o\<^sub>r\<^sub>y :: "Location \ int \ MTypes \ MemoryT \ String.literal \String.literal \String.literal" where +"dump\<^sub>M\<^sub>e\<^sub>m\<^sub>o\<^sub>r\<^sub>y loc x t mem ls str = iter (\i. dumprec\<^sub>M\<^sub>e\<^sub>m\<^sub>o\<^sub>r\<^sub>y ((hash loc (ShowL\<^sub>i\<^sub>n\<^sub>t i))) t mem (ls + STR ''['' + (ShowL\<^sub>i\<^sub>n\<^sub>t i + STR '']''))) str x" + +subsubsection\Storage\ + +datatype Data\<^sub>S\<^sub>t\<^sub>o\<^sub>r\<^sub>a\<^sub>g\<^sub>e = + SArray "Data\<^sub>S\<^sub>t\<^sub>o\<^sub>r\<^sub>a\<^sub>g\<^sub>e list" + | SMap "(String.literal \ Data\<^sub>S\<^sub>t\<^sub>o\<^sub>r\<^sub>a\<^sub>g\<^sub>e) list" + | SBool bool + | SInt int + | SAddress Address + +fun go\<^sub>S\<^sub>t\<^sub>o\<^sub>r\<^sub>a\<^sub>g\<^sub>e :: "Location \ (String.literal \ STypes) \ (String.literal \ STypes)" where + "go\<^sub>S\<^sub>t\<^sub>o\<^sub>r\<^sub>a\<^sub>g\<^sub>e l (s, STArray _ t) = (s + (STR ''['') + (convert l) + (STR '']''), t)" +| "go\<^sub>S\<^sub>t\<^sub>o\<^sub>r\<^sub>a\<^sub>g\<^sub>e l (s, STMap _ t) = (s + (STR ''['') + (convert l) + (STR '']''), t)" +| "go\<^sub>S\<^sub>t\<^sub>o\<^sub>r\<^sub>a\<^sub>g\<^sub>e l (s, STValue t) = (s + (STR ''['') + (convert l) + (STR '']''), STValue t)" + +fun dumpSingle\<^sub>S\<^sub>t\<^sub>o\<^sub>r\<^sub>a\<^sub>g\<^sub>e :: "StorageT \ String.literal \ STypes \ (Location \ Location) \ String.literal \ String.literal" where +"dumpSingle\<^sub>S\<^sub>t\<^sub>o\<^sub>r\<^sub>a\<^sub>g\<^sub>e sto id' tp (loc,l) str = + (case foldr go\<^sub>S\<^sub>t\<^sub>o\<^sub>r\<^sub>a\<^sub>g\<^sub>e (tolist loc) (str + id', tp) of + (s, STValue t) \ + (case sto $$ (loc + l) of + Some v \ s + (STR ''=='') + dump\<^sub>V\<^sub>a\<^sub>l\<^sub>u\<^sub>e\<^sub>t\<^sub>y\<^sub>p\<^sub>e\<^sub>s t v + | None \ FAILURE) + | _ \ FAILURE)" + +definition iterate where +"iterate loc t id' sto s l = dumpSingle\<^sub>S\<^sub>t\<^sub>o\<^sub>r\<^sub>a\<^sub>g\<^sub>e sto id' t (splitAt (length (String.explode l) - length (String.explode loc) - 1) l) s + (STR ''\'')" + +fun dump\<^sub>S\<^sub>t\<^sub>o\<^sub>r\<^sub>a\<^sub>g\<^sub>e :: "StorageT \ Location \ String.literal \ STypes \ String.literal \ String.literal" where + "dump\<^sub>S\<^sub>t\<^sub>o\<^sub>r\<^sub>a\<^sub>g\<^sub>e sto loc id' (STArray _ t) str = foldl (iterate loc t id' sto) str (locations_map loc sto)" +| "dump\<^sub>S\<^sub>t\<^sub>o\<^sub>r\<^sub>a\<^sub>g\<^sub>e sto loc id' (STMap _ t) str = foldl (iterate loc t id' sto) str (locations_map loc sto)" +| "dump\<^sub>S\<^sub>t\<^sub>o\<^sub>r\<^sub>a\<^sub>g\<^sub>e sto loc id' (STValue t) str = + (case sto $$ loc of + Some v \ str + id' + (STR ''=='') + dump\<^sub>V\<^sub>a\<^sub>l\<^sub>u\<^sub>e\<^sub>t\<^sub>y\<^sub>p\<^sub>e\<^sub>s t v + (STR ''\'') + | _ \ str)" + +fun loadRec\<^sub>S\<^sub>t\<^sub>o\<^sub>r\<^sub>a\<^sub>g\<^sub>e :: "Location \ Data\<^sub>S\<^sub>t\<^sub>o\<^sub>r\<^sub>a\<^sub>g\<^sub>e \ StorageT \ StorageT" and + iterateSA :: "Location \ StorageT \ nat \ Data\<^sub>S\<^sub>t\<^sub>o\<^sub>r\<^sub>a\<^sub>g\<^sub>e \ StorageT \ nat" and + iterateSM :: "Location \ String.literal \ Data\<^sub>S\<^sub>t\<^sub>o\<^sub>r\<^sub>a\<^sub>g\<^sub>e \ StorageT \ StorageT" where + "loadRec\<^sub>S\<^sub>t\<^sub>o\<^sub>r\<^sub>a\<^sub>g\<^sub>e loc (SArray dat) sto = fst (foldl (iterateSA loc) (sto,0) dat)" +| "loadRec\<^sub>S\<^sub>t\<^sub>o\<^sub>r\<^sub>a\<^sub>g\<^sub>e loc (SMap dat) sto = foldr (iterateSM loc) dat sto" +| "loadRec\<^sub>S\<^sub>t\<^sub>o\<^sub>r\<^sub>a\<^sub>g\<^sub>e loc (SBool b) sto = fmupd loc (ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l b) sto" +| "loadRec\<^sub>S\<^sub>t\<^sub>o\<^sub>r\<^sub>a\<^sub>g\<^sub>e loc (SInt i) sto = fmupd loc (ShowL\<^sub>i\<^sub>n\<^sub>t i) sto" +| "loadRec\<^sub>S\<^sub>t\<^sub>o\<^sub>r\<^sub>a\<^sub>g\<^sub>e loc (SAddress ad) sto = fmupd loc ad sto" +| "iterateSA loc (s', x) d = (loadRec\<^sub>S\<^sub>t\<^sub>o\<^sub>r\<^sub>a\<^sub>g\<^sub>e (hash loc (ShowL\<^sub>n\<^sub>a\<^sub>t x)) d s', Suc x)" +| "iterateSM loc (k, v) s' = loadRec\<^sub>S\<^sub>t\<^sub>o\<^sub>r\<^sub>a\<^sub>g\<^sub>e (hash loc k) v s'" + +subsubsection\Environment\ + +datatype Data\<^sub>E\<^sub>n\<^sub>v\<^sub>i\<^sub>r\<^sub>o\<^sub>n\<^sub>m\<^sub>e\<^sub>n\<^sub>t = + Memarr "Data\<^sub>M\<^sub>e\<^sub>m\<^sub>o\<^sub>r\<^sub>y list" | + CDarr "Data\<^sub>M\<^sub>e\<^sub>m\<^sub>o\<^sub>r\<^sub>y list" | + Stoarr "Data\<^sub>S\<^sub>t\<^sub>o\<^sub>r\<^sub>a\<^sub>g\<^sub>e list"| + Stomap "(String.literal \ Data\<^sub>S\<^sub>t\<^sub>o\<^sub>r\<^sub>a\<^sub>g\<^sub>e) list" | + Stackbool bool | + Stobool bool | + Stackint int | + Stoint int | + Stackaddr Address | + Stoaddr Address + +fun astore :: "Identifier \ Type \ Valuetype \ StorageT * Environment \ StorageT * Environment" + where "astore i t v (s, e) = (fmupd i v s, (updateEnv i t (Storeloc i) e))" + +fun loadsimple\<^sub>E\<^sub>n\<^sub>v\<^sub>i\<^sub>r\<^sub>o\<^sub>n\<^sub>m\<^sub>e\<^sub>n\<^sub>t :: "(Stack \ CalldataT \ MemoryT \ StorageT \ Environment) + \ (Identifier \ Type \ Data\<^sub>E\<^sub>n\<^sub>v\<^sub>i\<^sub>r\<^sub>o\<^sub>n\<^sub>m\<^sub>e\<^sub>n\<^sub>t) \ (Stack \ CalldataT \ MemoryT \ StorageT \ Environment)" + where +"loadsimple\<^sub>E\<^sub>n\<^sub>v\<^sub>i\<^sub>r\<^sub>o\<^sub>n\<^sub>m\<^sub>e\<^sub>n\<^sub>t (k, c, m, s, e) (id', tp, d) = (case d of + Stackbool b \ + let (k', e') = astack id' tp (KValue (ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l b)) (k, e) + in (k', c, m, s, e') + | Stobool b \ + let (s', e') = astore id' tp (ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l b) (s, e) + in (k, c, m, s', e') + | Stackint n \ + let (k', e') = astack id' tp (KValue (ShowL\<^sub>i\<^sub>n\<^sub>t n)) (k, e) + in (k', c, m, s, e') + | Stoint n \ + let (s', e') = astore id' tp (ShowL\<^sub>i\<^sub>n\<^sub>t n) (s, e) + in (k, c, m, s', e') + | Stackaddr ad \ + let (k', e') = astack id' tp (KValue ad) (k, e) + in (k', c, m, s, e') + | Stoaddr ad \ + let (s', e') = astore id' tp ad (s, e) + in (k, c, m, s', e') + | CDarr a \ + let l = ShowL\<^sub>n\<^sub>a\<^sub>t (toploc c); + c' = load\<^sub>M\<^sub>e\<^sub>m\<^sub>o\<^sub>r\<^sub>y a c; + (k', e') = astack id' tp (KCDptr l) (k, e) + in (k', c', m, s, e') + | Memarr a \ + let l = ShowL\<^sub>n\<^sub>a\<^sub>t (toploc m); + m' = load\<^sub>M\<^sub>e\<^sub>m\<^sub>o\<^sub>r\<^sub>y a m; + (k', e') = astack id' tp (KMemptr l) (k, e) + in (k', c, m', s, e') + | Stoarr a \ + let s' = loadRec\<^sub>S\<^sub>t\<^sub>o\<^sub>r\<^sub>a\<^sub>g\<^sub>e id' (SArray a) s; + e' = updateEnv id' tp (Storeloc id') e + in (k, c, m, s', e') + | Stomap mp \ + let s' = loadRec\<^sub>S\<^sub>t\<^sub>o\<^sub>r\<^sub>a\<^sub>g\<^sub>e id' (SMap mp) s; + e' = updateEnv id' tp (Storeloc id') e + in (k, c, m, s', e') +)" + +definition getValue\<^sub>E\<^sub>n\<^sub>v\<^sub>i\<^sub>r\<^sub>o\<^sub>n\<^sub>m\<^sub>e\<^sub>n\<^sub>t :: "Stack \ CalldataT \ MemoryT \ StorageT \ Environment \ Identifier \ String.literal \ String.literal" + where +"getValue\<^sub>E\<^sub>n\<^sub>v\<^sub>i\<^sub>r\<^sub>o\<^sub>n\<^sub>m\<^sub>e\<^sub>n\<^sub>t k c m s e i txt = (case fmlookup (denvalue e) i of + Some (tp, Stackloc l) \ (case accessStore l k of + Some (KValue v) \ (case tp of + Value t \ (txt + i + (STR ''=='') + dump\<^sub>V\<^sub>a\<^sub>l\<^sub>u\<^sub>e\<^sub>t\<^sub>y\<^sub>p\<^sub>e\<^sub>s t v + (STR ''\'')) + | _ \ FAILURE) + | Some (KCDptr p) \ (case tp of + Calldata (MTArray x t) \ dump\<^sub>M\<^sub>e\<^sub>m\<^sub>o\<^sub>r\<^sub>y p x t c i txt + | _ \ FAILURE) + | Some (KMemptr p) \ (case tp of + Memory (MTArray x t) \ dump\<^sub>M\<^sub>e\<^sub>m\<^sub>o\<^sub>r\<^sub>y p x t m i txt + | _ \ FAILURE) + | Some (KStoptr p) \ (case tp of + Storage t \ dump\<^sub>S\<^sub>t\<^sub>o\<^sub>r\<^sub>a\<^sub>g\<^sub>e s p i t txt + | _ \ FAILURE)) + | Some (Storage t, Storeloc l) \ dump\<^sub>S\<^sub>t\<^sub>o\<^sub>r\<^sub>a\<^sub>g\<^sub>e s l i t txt + | _ \ FAILURE +)" + +definition dump\<^sub>E\<^sub>n\<^sub>v\<^sub>i\<^sub>r\<^sub>o\<^sub>n\<^sub>m\<^sub>e\<^sub>n\<^sub>t :: "Stack \ CalldataT \ MemoryT \ StorageT \ Environment \ Identifier list \ String.literal" + where "dump\<^sub>E\<^sub>n\<^sub>v\<^sub>i\<^sub>r\<^sub>o\<^sub>n\<^sub>m\<^sub>e\<^sub>n\<^sub>t k c m s e sl = foldr (getValue\<^sub>E\<^sub>n\<^sub>v\<^sub>i\<^sub>r\<^sub>o\<^sub>n\<^sub>m\<^sub>e\<^sub>n\<^sub>t k c m s e) sl EMPTY" + +subsubsection\Accounts\ + +fun load\<^sub>A\<^sub>c\<^sub>c\<^sub>o\<^sub>u\<^sub>n\<^sub>t\<^sub>s :: "Accounts \ (Address \ Balance \ atype \ nat) list \ Accounts" where + "load\<^sub>A\<^sub>c\<^sub>c\<^sub>o\<^sub>u\<^sub>n\<^sub>t\<^sub>s acc [] = acc" + | "load\<^sub>A\<^sub>c\<^sub>c\<^sub>o\<^sub>u\<^sub>n\<^sub>t\<^sub>s acc ((ad, b, t, c)#as) = load\<^sub>A\<^sub>c\<^sub>c\<^sub>o\<^sub>u\<^sub>n\<^sub>t\<^sub>s (acc (ad:=\bal=b, type=Some t, contracts=c\)) as" + +fun dumpStorage::"StorageT \ (Identifier \ Member) \ String.literal" where + "dumpStorage s (i, Var x) = dump\<^sub>S\<^sub>t\<^sub>o\<^sub>r\<^sub>a\<^sub>g\<^sub>e s i i x EMPTY" +| "dumpStorage s (i, Function x) = FAILURE" +| "dumpStorage s (i, Method x) = FAILURE" + +fun dumpMembers :: "atype option \ Environment\<^sub>P \ StorageT \ String.literal" where + "dumpMembers None ep s = FAILURE" +| "dumpMembers (Some EOA) _ _ = STR ''EOA''" +| "dumpMembers (Some (Contract name)) ep s = + (case ep $$ name of + Some (ct, _) \ name + STR ''('' + (intersperse (STR '','') (map (dumpStorage s) (filter (is_Var \ snd) (sorted_list_of_fmap ct)))) + STR '')'' + | None \ FAILURE)" + +(* + The first parameter is just to guarantee termination +*) +fun dump\<^sub>A\<^sub>c\<^sub>c\<^sub>o\<^sub>u\<^sub>n\<^sub>t :: "nat \ Environment\<^sub>P \ State \ Address \ String.literal" + where "dump\<^sub>A\<^sub>c\<^sub>c\<^sub>o\<^sub>u\<^sub>n\<^sub>t 0 _ _ _ = FAILURE" + | "dump\<^sub>A\<^sub>c\<^sub>c\<^sub>o\<^sub>u\<^sub>n\<^sub>t (Suc c) ep st a = a + STR '': '' + + STR ''balance=='' + bal (accounts st a) + + STR '' - '' + dumpMembers (type (accounts st a)) ep (storage st a) + + iter (\x s. s + STR ''\'' + dump\<^sub>A\<^sub>c\<^sub>c\<^sub>o\<^sub>u\<^sub>n\<^sub>t c ep st (hash a (ShowL\<^sub>i\<^sub>n\<^sub>t x))) EMPTY (int (contracts (accounts st a)))" + +definition dump\<^sub>A\<^sub>c\<^sub>c\<^sub>o\<^sub>u\<^sub>n\<^sub>t\<^sub>s :: "Environment\<^sub>P \ State \ Address list \ String.literal" + where "dump\<^sub>A\<^sub>c\<^sub>c\<^sub>o\<^sub>u\<^sub>n\<^sub>t\<^sub>s ep st al = intersperse (STR ''\'') (map (dump\<^sub>A\<^sub>c\<^sub>c\<^sub>o\<^sub>u\<^sub>n\<^sub>t 1000 ep st) al)" + +definition init\<^sub>A\<^sub>c\<^sub>c\<^sub>o\<^sub>u\<^sub>n\<^sub>t::"(Address \ Balance \ atype \ nat) list => Accounts" where + "init\<^sub>A\<^sub>c\<^sub>c\<^sub>o\<^sub>u\<^sub>n\<^sub>t = load\<^sub>A\<^sub>c\<^sub>c\<^sub>o\<^sub>u\<^sub>n\<^sub>t\<^sub>s emptyAccount" + +type_synonym Data\<^sub>P = "(Identifier \ Member) list \ ((Identifier \ Type) list \ S) \ S" + +fun loadProc::"Identifier \ Data\<^sub>P \ Environment\<^sub>P \ Environment\<^sub>P" + where "loadProc i (xs, fb) = fmupd i (fmap_of_list xs, fb)" + +subsection\Test Setup\ + +definition(in statement_with_gas) eval::"Gas \ S \ Address \ Identifier \ Address \ Valuetype \ (Address \ Balance \ atype \ nat) list + \ (String.literal \ Type \ Data\<^sub>E\<^sub>n\<^sub>v\<^sub>i\<^sub>r\<^sub>o\<^sub>n\<^sub>m\<^sub>e\<^sub>n\<^sub>t) list + \ String.literal" + where "eval g stm addr name adest aval acc dat + = (let (k,c,m,s,e) = foldl loadsimple\<^sub>E\<^sub>n\<^sub>v\<^sub>i\<^sub>r\<^sub>o\<^sub>n\<^sub>m\<^sub>e\<^sub>n\<^sub>t (emptyStore, emptyStore, emptyStore, fmempty, emptyEnv addr name adest aval) dat; + a = init\<^sub>A\<^sub>c\<^sub>c\<^sub>o\<^sub>u\<^sub>n\<^sub>t acc; + s' = emptyStorage (addr := s); + z = \accounts=a,stack=k,memory=m,storage=s',gas=g\ + in ( + case (stmt stm e c z) of + Normal ((), z') \ (dump\<^sub>E\<^sub>n\<^sub>v\<^sub>i\<^sub>r\<^sub>o\<^sub>n\<^sub>m\<^sub>e\<^sub>n\<^sub>t (stack z') c (memory z') (storage z' addr) e (map (\ (a,b,c). a) dat)) + + (dump\<^sub>A\<^sub>c\<^sub>c\<^sub>o\<^sub>u\<^sub>n\<^sub>t\<^sub>s ep z' (map fst acc)) + | Exception Err \ STR ''Exception'' + | Exception Gas \ STR ''OutOfGas''))" + +global_interpretation soliditytest0: statement_with_gas costs_ex "fmap_of_list []" costs_min + defines stmt0 = "soliditytest0.stmt" + and lexp0 = soliditytest0.lexp + and expr0 = soliditytest0.expr + and ssel0 = soliditytest0.ssel + and rexp0 = soliditytest0.rexp + and msel0 = soliditytest0.msel + and load0 = soliditytest0.load + and eval0 = soliditytest0.eval + by unfold_locales auto + +lemma "eval0 1000 + SKIP + (STR ''089Be5381FcEa58aF334101414c04F993947C733'') + EMPTY + EMPTY + (STR ''0'') + [(STR ''089Be5381FcEa58aF334101414c04F993947C733'', STR ''100'', EOA, 0), (STR ''115f6e2F70210C14f7DB1AC69737a3CC78435d49'', STR ''100'', EOA, 0)] + [(STR ''v1'', (Value TBool, Stackbool True))] + = STR ''v1==true\089Be5381FcEa58aF334101414c04F993947C733: balance==100 - EOA\115f6e2F70210C14f7DB1AC69737a3CC78435d49: balance==100 - EOA''" + by(eval) + +lemma "eval0 1000 + SKIP + (STR ''089Be5381FcEa58aF334101414c04F993947C733'') + EMPTY + EMPTY + (STR ''0'') + [(STR ''089Be5381FcEa58aF334101414c04F993947C733'', STR ''100'', EOA, 0), (STR ''115f6e2F70210C14f7DB1AC69737a3CC78435d49'', STR ''100'', EOA, 0)] + [(STR ''v1'',(Memory (MTArray 5 (MTValue TBool)), Memarr [MBool True, MBool False, MBool True, MBool False, MBool True]))] + = STR ''v1[0]==true\v1[1]==false\v1[2]==true\v1[3]==false\v1[4]==true\089Be5381FcEa58aF334101414c04F993947C733: balance==100 - EOA\115f6e2F70210C14f7DB1AC69737a3CC78435d49: balance==100 - EOA''" + by(eval) + +lemma "eval0 1000 + (ITE FALSE (ASSIGN (Id (STR ''x'')) TRUE) (ASSIGN (Id (STR ''y'')) TRUE)) + (STR ''089Be5381FcEa58aF334101414c04F993947C733'') + EMPTY + EMPTY + (STR ''0'') + [(STR ''089Be5381FcEa58aF334101414c04F993947C733'', STR ''100'', EOA, 0), (STR ''115f6e2F70210C14f7DB1AC69737a3CC78435d49'', STR ''100'', EOA, 0)] + [(STR ''x'', (Value TBool, Stackbool False)), (STR ''y'', (Value TBool, Stackbool False))] + = STR ''y==true\x==false\089Be5381FcEa58aF334101414c04F993947C733: balance==100 - EOA\115f6e2F70210C14f7DB1AC69737a3CC78435d49: balance==100 - EOA''" + by(eval) + +lemma "eval0 1000 + (BLOCK ((STR ''v2'', Value TBool), None) (ASSIGN (Id (STR ''v1'')) (LVAL (Id (STR ''v2''))))) + (STR ''089Be5381FcEa58aF334101414c04F993947C733'') + EMPTY + EMPTY + (STR ''0'') + [(STR ''089Be5381FcEa58aF334101414c04F993947C733'', STR ''100'', EOA, 0), (STR ''115f6e2F70210C14f7DB1AC69737a3CC78435d49'', STR ''100'', EOA, 0)] + [(STR ''v1'', (Value TBool, Stackbool True))] + = STR ''v1==false\089Be5381FcEa58aF334101414c04F993947C733: balance==100 - EOA\115f6e2F70210C14f7DB1AC69737a3CC78435d49: balance==100 - EOA''" + by(eval) + +lemma "eval0 1000 + (ASSIGN (Id (STR ''a_s120_21_m8'')) (LVAL (Id (STR ''a_s120_21_s8'')))) + (STR ''089Be5381FcEa58aF334101414c04F993947C733'') + EMPTY + EMPTY + (STR ''0'') + [(STR ''089Be5381FcEa58aF334101414c04F993947C733'', STR ''100'', EOA, 0)] + [((STR ''a_s120_21_s8''), Storage (STArray 1 (STArray 2 (STValue (TSInt 120)))), Stoarr [SArray [SInt 347104507864064359095275590289383142, SInt 565831699297331399489670920129618233]]), + ((STR ''a_s120_21_m8''), Memory (MTArray 1 (MTArray 2 (MTValue (TSInt 120)))), Memarr [MArray [MInt (290845675805142398428016622247257774), MInt ((-96834026877269277170645294669272226))]])] + = STR ''a_s120_21_m8[0][0]==347104507864064359095275590289383142\a_s120_21_m8[0][1]==565831699297331399489670920129618233\a_s120_21_s8[0][0]==347104507864064359095275590289383142\a_s120_21_s8[0][1]==565831699297331399489670920129618233\089Be5381FcEa58aF334101414c04F993947C733: balance==100 - EOA''" + by(eval) + +lemma "eval0 1000 + (ASSIGN (Ref (STR ''a_s8_32_m0'') [UINT 8 1]) (LVAL (Ref (STR ''a_s8_31_s7'') [UINT 8 0]))) + (STR ''089Be5381FcEa58aF334101414c04F993947C733'') + EMPTY + EMPTY + (STR ''0'') + [(STR ''089Be5381FcEa58aF334101414c04F993947C733'', STR ''100'', EOA, 0)] + [(STR ''a_s8_31_s7'', (Storage (STArray 1 (STArray 3 (STValue (TSInt 8)))), Stoarr [SArray [SInt ((98)), SInt ((-23)), SInt (36)]])), + (STR ''a_s8_32_m0'', (Memory (MTArray 2 (MTArray 3 (MTValue (TSInt 8)))), Memarr [MArray [MInt ((-64)), MInt ((39)), MInt ((-125))], MArray [MInt ((-32)), MInt ((-82)), MInt ((-105))]]))] + = STR ''a_s8_32_m0[0][0]==-64\a_s8_32_m0[0][1]==39\a_s8_32_m0[0][2]==-125\a_s8_32_m0[1][0]==98\a_s8_32_m0[1][1]==-23\a_s8_32_m0[1][2]==36\a_s8_31_s7[0][0]==98\a_s8_31_s7[0][1]==-23\a_s8_31_s7[0][2]==36\089Be5381FcEa58aF334101414c04F993947C733: balance==100 - EOA''" + by(eval) + +lemma "eval0 1000 + SKIP + (STR ''089Be5381FcEa58aF334101414c04F993947C733'') + EMPTY + EMPTY + (STR ''0'') + [(STR ''089Be5381FcEa58aF334101414c04F993947C733'', STR ''100'', EOA, 0), (STR ''115f6e2F70210C14f7DB1AC69737a3CC78435d49'', STR ''100'', EOA, 0)] + [(STR ''v1'', (Storage (STMap (TUInt 32) (STValue (TUInt 8))), Stomap [(STR ''2129136830'', SInt (247))]))] + = STR ''v1[2129136830]==247\089Be5381FcEa58aF334101414c04F993947C733: balance==100 - EOA\115f6e2F70210C14f7DB1AC69737a3CC78435d49: balance==100 - EOA''" + by(eval) + +definition "testenv1 \ loadProc (STR ''mycontract'') + ([(STR ''v1'', Var (STValue TBool)), + (STR ''m1'', Method ([], True, (ASSIGN (Id (STR ''v1'')) FALSE)))], + ([], SKIP), + SKIP) + fmempty" + +global_interpretation soliditytest1: statement_with_gas costs_ex testenv1 costs_min + defines stmt1 = "soliditytest1.stmt" + and lexp1 = soliditytest1.lexp + and expr1 = soliditytest1.expr + and ssel1 = soliditytest1.ssel + and rexp1 = soliditytest1.rexp + and msel1 = soliditytest1.msel + and load1 = soliditytest1.load + and eval1 = soliditytest1.eval + by unfold_locales auto + +lemma "eval1 1000 + (EXTERNAL (ADDRESS (STR ''myaddr'')) (STR ''m1'') [] (UINT 256 0)) + (STR ''local'') + EMPTY + (STR ''mycontract'') + (STR ''0'') + [(STR ''local'', STR ''100'', EOA, 0), (STR ''myaddr'', STR ''100'', Contract (STR ''mycontract''), 0)] + [] + = STR ''local: balance==100 - EOA\myaddr: balance==100 - mycontract(v1==false\)''" + by (eval) + +lemma "eval1 1000 + (NEW (STR ''mycontract'') [] (UINT 256 10)) + (STR ''local'') + EMPTY + (STR ''mycontract'') + (STR ''0'') + [(STR ''local'', STR ''100'', EOA, 0), (STR ''myaddr'', STR ''100'', Contract (STR ''mycontract''), 0)] + [] + = STR ''local: balance==90 - EOA\0.local: balance==10 - mycontract()\myaddr: balance==100 - mycontract()''" + by eval + +lemma "eval1 1000 + ( + COMP + (NEW (STR ''mycontract'') [] (UINT 256 10)) + (EXTERNAL CONTRACTS (STR ''m1'') [] (UINT 256 0)) + ) + (STR ''local'') + EMPTY + (STR ''mycontract'') + (STR ''0'') + [(STR ''local'', STR ''100'', EOA, 0), (STR ''myaddr'', STR ''100'', Contract (STR ''mycontract''), 0)] + [] + = STR ''local: balance==90 - EOA\0.local: balance==10 - mycontract(v1==false\)\myaddr: balance==100 - mycontract()''" + by eval + +definition "testenv2 \ loadProc (STR ''mycontract'') + ([(STR ''m1'', Function ([], False, UINT 8 5))], + ([], SKIP), + SKIP) + fmempty" + +global_interpretation soliditytest2: statement_with_gas costs_ex testenv2 costs_min + defines stmt2 = "soliditytest2.stmt" + and lexp2 = soliditytest2.lexp + and expr2 = soliditytest2.expr + and ssel2 = soliditytest2.ssel + and rexp2 = soliditytest2.rexp + and msel2 = soliditytest2.msel + and load2 = soliditytest2.load + and eval2 = soliditytest2.eval + by unfold_locales auto + +lemma "eval2 1000 + (ASSIGN (Id (STR ''v1'')) (CALL (STR ''m1'') [])) + (STR ''myaddr'') + (STR ''mycontract'') + EMPTY + (STR ''0'') + [(STR ''myaddr'', STR ''100'', EOA, 0)] + [(STR ''v1'', (Value (TUInt 8), Stackint 0))] + = STR ''v1==5\myaddr: balance==100 - EOA''" + by (eval) + +definition "testenv3 \ loadProc (STR ''mycontract'') + ([(STR ''m1'', + Function ([(STR ''v2'', Value (TSInt 8)), (STR ''v3'', Value (TSInt 8))], + False, + PLUS (LVAL (Id (STR ''v2''))) (LVAL (Id (STR ''v3'')))))], + ([], SKIP), + SKIP) + fmempty" + +global_interpretation soliditytest3: statement_with_gas costs_ex testenv3 costs_min + defines stmt3 = "soliditytest3.stmt" + and lexp3 = soliditytest3.lexp + and expr3 = soliditytest3.expr + and ssel3 = soliditytest3.ssel + and rexp3 = soliditytest3.rexp + and msel3 = soliditytest3.msel + and load3 = soliditytest3.load + and eval3 = soliditytest3.eval + by unfold_locales auto + +lemma "eval3 1000 + (ASSIGN (Id (STR ''v1'')) (CALL (STR ''m1'') [E.INT 8 3, E.INT 8 4])) + (STR ''myaddr'') + (STR ''mycontract'') + EMPTY + (STR ''0'') + [(STR ''myaddr'', STR ''100'', EOA, 0),(STR ''mya'', STR ''100'', EOA, 0)] + [(STR ''v1'', (Value (TSInt 8), Stackint 0))] + = STR ''v1==7\myaddr: balance==100 - EOA\mya: balance==100 - EOA''" + by (eval) + +definition "testenv4 \ loadProc (STR ''mycontract'') + ([(STR ''m1'', Function ([(STR ''v2'', Value (TSInt 8)), (STR ''v3'', Value (TSInt 8))], True, PLUS (LVAL (Id (STR ''v2''))) (LVAL (Id (STR ''v3'')))))], + ([], SKIP), + SKIP) + fmempty" + +global_interpretation soliditytest4: statement_with_gas costs_ex testenv4 costs_min + defines stmt4 = "soliditytest4.stmt" + and lexp4 = soliditytest4.lexp + and expr4 = soliditytest4.expr + and ssel4 = soliditytest4.ssel + and rexp4 = soliditytest4.rexp + and msel4 = soliditytest4.msel + and load4 = soliditytest4.load + and eval4 = soliditytest4.eval + by unfold_locales auto + +lemma "eval4 1000 + (ASSIGN (Id (STR ''v1'')) (ECALL (ADDRESS (STR ''extaddr'')) (STR ''m1'') [E.INT 8 3, E.INT 8 4])) + (STR ''myaddr'') + EMPTY + EMPTY + (STR ''0'') + [(STR ''myaddr'', STR ''100'', EOA, 0), (STR ''extaddr'', STR ''100'', Contract (STR ''mycontract''), 0)] + [(STR ''v1'', (Value (TSInt 8), Stackint 0))] + = STR ''v1==7\myaddr: balance==100 - EOA\extaddr: balance==100 - mycontract()''" + by (eval) + +definition "testenv5 \ loadProc (STR ''mycontract'') + ([], ([], SKIP), SKIP) + fmempty" + +global_interpretation soliditytest5: statement_with_gas costs_ex testenv5 costs_min + defines stmt5 = "soliditytest5.stmt" + and lexp5 = soliditytest5.lexp + and expr5 = soliditytest5.expr + and ssel5 = soliditytest5.ssel + and rexp5 = soliditytest5.rexp + and msel5 = soliditytest5.msel + and load5 = soliditytest5.load + and eval5 = soliditytest5.eval + by unfold_locales auto + +lemma "eval5 1000 + (TRANSFER (ADDRESS (STR ''myaddr'')) (UINT 256 10)) + (STR ''089Be5381FcEa58aF334101414c04F993947C733'') + EMPTY + EMPTY + (STR ''0'') + [(STR ''089Be5381FcEa58aF334101414c04F993947C733'', STR ''100'', EOA, 0), (STR ''myaddr'', STR ''100'', Contract (STR ''mycontract''), 0)] + [] + = STR ''089Be5381FcEa58aF334101414c04F993947C733: balance==90 - EOA\myaddr: balance==110 - mycontract()''" + by (eval) + +definition "testenv6 \ loadProc (STR ''Receiver'') + ([(STR ''hello'', Var (STValue TBool))], + ([], SKIP), + ASSIGN (Id (STR ''hello'')) TRUE) + fmempty" + +global_interpretation soliditytest6: statement_with_gas costs_ex testenv6 costs_min + defines stmt6 = "soliditytest6.stmt" + and lexp6 = soliditytest6.lexp + and expr6 = soliditytest6.expr + and ssel6 = soliditytest6.ssel + and rexp6 = soliditytest6.rexp + and msel6 = soliditytest6.msel + and load6 = soliditytest6.load + and eval6 = soliditytest6.eval + by unfold_locales auto + +lemma "eval6 1000 + (TRANSFER (ADDRESS (STR ''ReceiverAd'')) (UINT 256 10)) + (STR ''SenderAd'') + EMPTY + EMPTY + (STR ''0'') + [(STR ''ReceiverAd'', STR ''100'', Contract (STR ''Receiver''), 0), (STR ''SenderAd'', STR ''100'', EOA, 0)] + [] + = STR ''ReceiverAd: balance==110 - Receiver(hello==true\)\SenderAd: balance==90 - EOA''" + by (eval) + +definition "testenv7 \ loadProc (STR ''mycontract'') + ([], ([], SKIP), SKIP) + fmempty" + +global_interpretation soliditytest7: statement_with_gas costs_ex testenv7 costs_min + defines stmt7 = "soliditytest7.stmt" + and lexp7 = soliditytest7.lexp + and expr7 = soliditytest7.expr + and ssel7 = soliditytest7.ssel + and rexp7 = soliditytest7.rexp + and msel7 = soliditytest7.msel + and load7 = soliditytest7.load + and eval7 = soliditytest7.eval + by unfold_locales auto + +lemma "eval7 1000 + (COMP(COMP(((ASSIGN (Id (STR ''x'')) (E.UINT 8 0))))(TRANSFER (ADDRESS (STR ''myaddr'')) (UINT 256 5)))(SKIP)) + (STR ''089Be5381FcEa58aF334101414c04F993947C733'') + EMPTY + EMPTY + (STR ''0'') + [(STR ''089Be5381FcEa58aF334101414c04F993947C733'', STR ''100'', EOA, 0), (STR ''myaddr'', STR ''100'', Contract (STR ''mycontract''),0)] + [(STR ''x'', (Value (TUInt 8), Stackint 9))] + = STR ''x==0\089Be5381FcEa58aF334101414c04F993947C733: balance==95 - EOA\myaddr: balance==105 - mycontract()''" + by (eval) + +subsection\The Final Code Export\ + +consts ReadL\<^sub>S :: "String.literal \ S" +consts ReadL\<^sub>a\<^sub>c\<^sub>c :: "String.literal \ (String.literal \ String.literal \ atype \ nat) list" +consts ReadL\<^sub>d\<^sub>a\<^sub>t :: "String.literal \ (String.literal \ Type \ Data\<^sub>E\<^sub>n\<^sub>v\<^sub>i\<^sub>r\<^sub>o\<^sub>n\<^sub>m\<^sub>e\<^sub>n\<^sub>t) list" +consts ReadL\<^sub>P :: "String.literal \ Data\<^sub>P list" + +code_printing + constant ReadL\<^sub>S \ (Haskell) "Prelude.read" + | constant ReadL\<^sub>a\<^sub>c\<^sub>c \ (Haskell) "Prelude.read" + | constant ReadL\<^sub>d\<^sub>a\<^sub>t \ (Haskell) "Prelude.read" + | constant ReadL\<^sub>P \ (Haskell) "Prelude.read" + +fun main_stub :: "String.literal list \ (int \ String.literal)" + where + "main_stub [stm, saddr, name, raddr, val, acc, dat] + = (0, eval0 1000 (ReadL\<^sub>S stm) saddr name raddr val (ReadL\<^sub>a\<^sub>c\<^sub>c acc) (ReadL\<^sub>d\<^sub>a\<^sub>t dat))" + | "main_stub _ = (2, + STR ''solidity-evaluator [credit] \"Statement\" \"ContractAddress\" \"OriginAddress\" \"Value\"\'' + + STR '' \"(Address * Balance) list\" \"(Address * ((Identifier * Member) list) * Statement)\" \"(Variable * Type * Value) list\"\'' + + STR ''\'')" + +generate_file "code/solidity-evaluator/app/Main.hs" = \ +module Main where +import System.Environment +import Solidity_Evaluator +import Prelude + +main :: IO () +main = do + args <- getArgs + Prelude.putStr(snd $ Solidity_Evaluator.main_stub args) +\ + +export_generated_files _ + +export_code eval0 SKIP main_stub + in Haskell module_name "Solidity_Evaluator" file_prefix "solidity-evaluator/src" (string_classes) + +subsection\Demonstrating the Symbolic Execution of Solidity\ + +abbreviation P1::S + where "P1 \ COMP (ASSIGN (Id (STR ''sa'')) (LVAL (Id (STR ''ma'')))) + (ASSIGN (Ref (STR ''sa'') [UINT (8::nat) 0]) TRUE)" +abbreviation myenv::Environment + where "myenv \ updateEnv (STR ''ma'') (Memory (MTArray 1 (MTValue TBool))) (Stackloc (STR ''1'')) + (updateEnv (STR ''sa'') (Storage (STArray 1 (STValue TBool))) (Storeloc (STR ''1'')) + (emptyEnv (STR ''ad'') EMPTY (STR ''ad'') (STR ''0'')))" + +abbreviation mystack::Stack + where "mystack \ updateStore (STR ''1'') (KMemptr (STR ''1 '')) emptyStore" + +abbreviation mystore::StorageT + where "mystore \ fmempty" + +abbreviation mymemory::MemoryT + where "mymemory \ updateStore (STR ''0.1'') (MValue (STR ''false'')) emptyStore" + +abbreviation mystorage::StorageT + where "mystorage \ fmupd (STR ''0.1'') (STR ''True'') fmempty" + +(* + FIXME: + +lemma "( stmt P1 myenv emptyStore \accounts=emptyAccount, stack=mystack, memory=mymemory, storage=fmupd (STR ''ad'') mystorage fmempty, gas=1000\ ) = + (Normal ((), \accounts=emptyAccount, stack=mystack, memory=mymemory, storage=fmupd (STR ''ad'') mystorage fmempty, gas=1000\ ))" + by(solidity_symbex) +*) +end diff --git a/thys/Solidity/Solidity_Main.thy b/thys/Solidity/Solidity_Main.thy --- a/thys/Solidity/Solidity_Main.thy +++ b/thys/Solidity/Solidity_Main.thy @@ -1,17 +1,17 @@ -section\ The Main Entry Point\ -theory - Solidity_Main -imports - Valuetypes - Storage - Environment - Statements -begin -text\ - This theory is the main entry point into the session Solidity, i.e., it serves the same purpose - as @{theory "Main"} for the session @{session "HOL"}. - - It is based on Solidity v0.5.16~\url{https://docs.soliditylang.org/en/v0.5.16/index.html} -\ - -end +section\ The Main Entry Point\ +theory + Solidity_Main +imports + Valuetypes + Storage + Environment + Statements +begin +text\ + This theory is the main entry point into the session Solidity, i.e., it serves the same purpose + as @{theory "Main"} for the session @{session "HOL"}. + + It is based on Solidity v0.5.16~\url{https://docs.soliditylang.org/en/v0.5.16/index.html} +\ + +end diff --git a/thys/Solidity/Solidity_Symbex.thy b/thys/Solidity/Solidity_Symbex.thy --- a/thys/Solidity/Solidity_Symbex.thy +++ b/thys/Solidity/Solidity_Symbex.thy @@ -1,29 +1,29 @@ -section\Towards a Setup for Symbolic Evaluation of Solidity\ -text\ - In this chapter, we lay out the foundations for a tactic for executing - Solidity statements and expressions symbolically. -\ -theory Solidity_Symbex -imports - Main - "HOL-Eisbach.Eisbach" -begin - -lemma string_literal_cat: "a+b = String.implode ((String.explode a) @ (String.explode b))" - by (metis String.implode_explode_eq plus_literal.rep_eq) - - -lemma string_literal_conv: "(map String.ascii_of y = y) \ (x = String.implode y) = (String.explode x = y) " - by auto - -lemmas string_literal_opt = Literal.rep_eq zero_literal.rep_eq plus_literal.rep_eq - string_literal_cat string_literal_conv - -named_theorems solidity_symbex -method solidity_symbex declares solidity_symbex = - ((simp add:solidity_symbex cong:unit.case), (simp add:string_literal_opt)?; (code_simp,simp?)+) - -declare Let_def [solidity_symbex] - o_def [solidity_symbex] - -end +section\Towards a Setup for Symbolic Evaluation of Solidity\ +text\ + In this chapter, we lay out the foundations for a tactic for executing + Solidity statements and expressions symbolically. +\ +theory Solidity_Symbex +imports + Main + "HOL-Eisbach.Eisbach" +begin + +lemma string_literal_cat: "a+b = String.implode ((String.explode a) @ (String.explode b))" + by (metis String.implode_explode_eq plus_literal.rep_eq) + + +lemma string_literal_conv: "(map String.ascii_of y = y) \ (x = String.implode y) = (String.explode x = y) " + by auto + +lemmas string_literal_opt = Literal.rep_eq zero_literal.rep_eq plus_literal.rep_eq + string_literal_cat string_literal_conv + +named_theorems solidity_symbex +method solidity_symbex declares solidity_symbex = + ((simp add:solidity_symbex cong:unit.case), (simp add:string_literal_opt)?; (code_simp,simp?)+) + +declare Let_def [solidity_symbex] + o_def [solidity_symbex] + +end diff --git a/thys/Solidity/StateMonad.thy b/thys/Solidity/StateMonad.thy --- a/thys/Solidity/StateMonad.thy +++ b/thys/Solidity/StateMonad.thy @@ -1,251 +1,435 @@ -theory StateMonad -imports Main "HOL-Library.Monad_Syntax" -begin - -section "State Monad with Exceptions" - -datatype ('n, 'e) result = Normal 'n | Exception 'e - -type_synonym ('a, 'e, 's) state_monad = "'s \ ('a \ 's, 'e) result" - -lemma result_cases[cases type: result]: - fixes x :: "('a \ 's, 'e) result" - obtains (n) a s where "x = Normal (a, s)" - | (e) e where "x = Exception e" -proof (cases x) - case (Normal n) - then show ?thesis - by (metis n prod.swap_def swap_swap) - next - case (Exception e) - then show ?thesis .. -qed - -subsection \Fundamental Definitions\ - -fun - return :: "'a \ ('a, 'e, 's) state_monad" where - "return a s = Normal (a, s)" - -fun - throw :: "'e \ ('a, 'e, 's) state_monad" where - "throw e s = Exception e" - -fun - bind :: "('a, 'e, 's) state_monad \ ('a \ ('b, 'e, 's) state_monad) \ - ('b, 'e, 's) state_monad" (infixl ">>=" 60) - where - "bind f g s = (case f s of - Normal (a, s') \ g a s' - | Exception e \ Exception e)" - -adhoc_overloading Monad_Syntax.bind bind - -lemma throw_left[simp]: "throw x \ y = throw x" by auto - -subsection \The Monad Laws\ - -text \@{term return} is absorbed at the left of a @{term bind}, - applying the return value directly:\ -lemma return_bind [simp]: "(return x \ f) = f x" - by auto - -text \@{term return} is absorbed on the right of a @{term bind}\ -lemma bind_return [simp]: "(m \ return) = m" -proof - - have "\s. bind m return s = m s" - proof - fix s - show "bind m return s = m s" - proof (cases "m s" rule: result_cases) - case (n a s) - then show ?thesis by auto - next - case (e e) - then show ?thesis by auto - qed - qed - thus ?thesis by auto -qed - -text \@{term bind} is associative\ -lemma bind_assoc: - fixes m :: "('a,'e,'s) state_monad" - fixes f :: "'a \ ('b,'e,'s) state_monad" - fixes g :: "'b \ ('c,'e,'s) state_monad" - shows "(m \ f) \ g = m \ (\x. f x\ g)" -proof - fix s - show "bind (bind m f) g s = bind m (\x. bind (f x) g) s" - by (cases "m s" rule: result_cases, simp+) -qed - -subsection \Basic Conguruence Rules\ - -(* - This lemma is needed for termination proofs. - - Function sub1 , for example, requires it. -*) -lemma monad_cong[fundef_cong]: - fixes m1 m2 m3 m4 - assumes "m1 s = m2 s" - and "\v s'. m2 s = Normal (v, s') \ m3 v s' = m4 v s'" - shows "(bind m1 m3) s = (bind m2 m4) s" - apply(insert assms, cases "m1 s") - apply (metis StateMonad.bind.simps old.prod.case result.simps(5)) - by (metis bind.elims result.simps(6)) - -(* - Lemma bind_case_nat_cong is required if a bind operand is a case analysis over nat. - - Function sub2, for example, requires it. -*) -lemma bind_case_nat_cong [fundef_cong]: - assumes "x = x'" and "\a. x = Suc a \ f a h = f' a h" - shows "(case x of Suc a \ f a | 0 \ g) h = (case x' of Suc a \ f' a | 0 \ g) h" - by (metis assms(1) assms(2) old.nat.exhaust old.nat.simps(4) old.nat.simps(5)) - -lemma if_cong[fundef_cong]: - assumes "b = b'" - and "b' \ m1 s = m1' s" - and "\ b' \ m2 s = m2' s" - shows "(if b then m1 else m2) s = (if b' then m1' else m2') s" - using assms(1) assms(2) assms(3) by auto - -lemma bind_case_pair_cong [fundef_cong]: - assumes "x = x'" and "\a b. x = (a,b) \ f a b s = f' a b s" - shows "(case x of (a,b) \ f a b) s = (case x' of (a,b) \ f' a b) s" - by (simp add: assms(1) assms(2) prod.case_eq_if) - -lemma bind_case_let_cong [fundef_cong]: - assumes "M = N" - and "(\x. x = N \ f x s = g x s)" - shows "(Let M f) s = (Let N g) s" - by (simp add: assms(1) assms(2)) - -lemma bind_case_some_cong [fundef_cong]: - assumes "x = x'" and "\a. x = Some a \ f a s = f' a s" and "x = None \ g s = g' s" - shows "(case x of Some a \ f a | None \ g) s = (case x' of Some a \ f' a | None \ g') s" - by (simp add: assms(1) assms(2) assms(3) option.case_eq_if) - -subsection \Other functions\ - -text \ - The basic accessor functions of the state monad. \get\ returns - the current state as result, does not fail, and does not change the state. - \put s\ returns unit, changes the current state to \s\ and does not fail. -\ -fun get :: "('s, 'e, 's) state_monad" where - "get s = Normal (s, s)" - -fun put :: "'s \ (unit, 'e, 's) state_monad" where - "put s _ = Normal ((), s)" - -text \Apply a function to the current state and return the result -without changing the state.\ -fun - applyf :: "('s \ 'a) \ ('a, 'e, 's) state_monad" where - "applyf f = get \ (\s. return (f s))" - -text \Modify the current state using the function passed in.\ -fun - modify :: "('s \ 's) \ (unit, 'e, 's) state_monad" -where "modify f = get \ (\s::'s. put (f s))" - -text \ - Perform a test on the current state, performing the left monad if - the result is true or the right monad if the result is false. -\ -fun - condition :: "('s \ bool) \ ('a, 'e, 's) state_monad \ ('a, 'e, 's) state_monad \ ('a, 'e, 's) state_monad" -where - "condition P L R s = (if (P s) then (L s) else (R s))" - -notation (output) - condition ("(condition (_)// (_)// (_))" [1000,1000,1000] 1000) - -lemma condition_cong[fundef_cong]: - assumes "b s = b' s" - and "b' s \ m1 s = m1' s" - and "\s'. s' = s \ \ b' s' \ m2 s' = m2' s'" - shows "(condition b m1 m2) s = (condition b' m1' m2') s" - by (simp add: assms(1) assms(2) assms(3)) - -fun - assert :: "'e \ ('s \ bool) \ ('a, 'e, 's) state_monad \ ('a, 'e, 's) state_monad" where - "assert x t m = condition t (throw x) m" - -notation (output) - assert ("(assert (_)// (_)// (_))" [1000,1000,1000] 1000) - -lemma assert_cong[fundef_cong]: - assumes "b s = b' s" - and "\ b' s \ m s = m' s" - shows "(assert x b m) s = (assert x b' m') s" - by (simp add: assms(1) assms(2)) - -subsection \Some basic examples\ - -lemma "do { - x \ return 1; - return (2::nat); - return x - } = - return 1 \ (\x. return (2::nat) \ (\_. (return x)))" .. - -lemma "do { - x \ return 1; - return 2; - return x - } = return 1" - by auto - -fun sub1 :: "(unit, nat, nat) state_monad" where - "sub1 0 = put 0 0" - | "sub1 (Suc n) = (do { - x \ get; - put x; - sub1 - }) n" - -fun sub2 :: "(unit, nat, nat) state_monad" where - "sub2 s = - (do { - n \ get; - (case n of - 0 \ put 0 - | Suc n' \ (do { - put n'; - sub2 - })) - }) s" - -fun sub3 :: "(unit, nat, nat) state_monad" where - "sub3 s = - condition (\n. n=0) - (return ()) - (do { - n \ get; - put (n - 1); - sub3 - }) s" - -fun sub4 :: "(unit, nat, nat) state_monad" where - "sub4 s = - assert (0) (\n. n=0) - (do { - n \ get; - put (n - 1); - sub4 - }) s" - -fun sub5 :: "(unit, nat, (nat*nat)) state_monad" where - "sub5 s = - assert (0) (\n. fst n=0) - (do { - (n,m) \ get; - put (n - 1,m); - sub5 - }) s" -end \ No newline at end of file +theory StateMonad +imports Main "HOL-Library.Monad_Syntax" Utils +begin + +section "State Monad with Exceptions" + +datatype ('n, 'e) result = Normal (normal: 'n) | Exception (exception: 'e) + +type_synonym ('a, 'e, 's) state_monad = "'s \ ('a \ 's, 'e) result" + +lemma result_cases[cases type: result]: + fixes x :: "('a \ 's, 'e) result" + obtains (n) a s where "x = Normal (a, s)" + | (e) e where "x = Exception e" +proof (cases x) + case (Normal n) + then show ?thesis + by (metis n prod.swap_def swap_swap) + next + case (Exception e) + then show ?thesis .. +qed + +subsection \Fundamental Definitions\ + +fun return :: "'a \ ('a, 'e, 's) state_monad" +where "return a s = Normal (a, s)" + +fun throw :: "'e \ ('a, 'e, 's) state_monad" +where "throw e s = Exception e" + +fun bind :: "('a, 'e, 's) state_monad \ ('a \ ('b, 'e, 's) state_monad) \ ('b, 'e, 's) state_monad" (infixl ">>=" 60) +where "bind f g s = (case f s of + Normal (a, s') \ g a s' + | Exception e \ Exception e)" + +adhoc_overloading Monad_Syntax.bind bind + +lemma throw_left[simp]: "throw x \ y = throw x" by auto + +subsection \The Monad Laws\ + +text \@{term return} is absorbed at the left of a @{term bind}, + applying the return value directly:\ +lemma return_bind [simp]: "(return x \ f) = f x" + by auto + +text \@{term return} is absorbed on the right of a @{term bind}\ +lemma bind_return [simp]: "(m \ return) = m" +proof - + have "\s. bind m return s = m s" + proof + fix s + show "bind m return s = m s" + proof (cases "m s" rule: result_cases) + case (n a s) + then show ?thesis by auto + next + case (e e) + then show ?thesis by auto + qed + qed + thus ?thesis by auto +qed + +text \@{term bind} is associative\ +lemma bind_assoc: + fixes m :: "('a,'e,'s) state_monad" + fixes f :: "'a \ ('b,'e,'s) state_monad" + fixes g :: "'b \ ('c,'e,'s) state_monad" + shows "(m \ f) \ g = m \ (\x. f x\ g)" +proof + fix s + show "bind (bind m f) g s = bind m (\x. bind (f x) g) s" + by (cases "m s" rule: result_cases, simp+) +qed + +subsection \Basic Conguruence Rules\ + +(* + This lemma is needed for termination proofs. + + Function sub1 , for example, requires it. +*) +lemma monad_cong[fundef_cong]: + fixes m1 m2 m3 m4 + assumes "m1 s = m2 s" + and "\v s'. m2 s = Normal (v, s') \ m3 v s' = m4 v s'" + shows "(bind m1 m3) s = (bind m2 m4) s" + apply(insert assms, cases "m1 s") + apply (metis StateMonad.bind.simps old.prod.case result.simps(5)) + by (metis bind.elims result.simps(6)) + +(* + Lemma bind_case_nat_cong is required if a bind operand is a case analysis over nat. + + Function sub2, for example, requires it. +*) +lemma bind_case_nat_cong [fundef_cong]: + assumes "x = x'" and "\a. x = Suc a \ f a h = f' a h" + shows "(case x of Suc a \ f a | 0 \ g) h = (case x' of Suc a \ f' a | 0 \ g) h" + by (metis assms(1) assms(2) old.nat.exhaust old.nat.simps(4) old.nat.simps(5)) + +lemma if_cong[fundef_cong]: + assumes "b = b'" + and "b' \ m1 s = m1' s" + and "\ b' \ m2 s = m2' s" + shows "(if b then m1 else m2) s = (if b' then m1' else m2') s" + using assms(1) assms(2) assms(3) by auto + +lemma bind_case_pair_cong [fundef_cong]: + assumes "x = x'" and "\a b. x = (a,b) \ f a b s = f' a b s" + shows "(case x of (a,b) \ f a b) s = (case x' of (a,b) \ f' a b) s" + by (simp add: assms(1) assms(2) prod.case_eq_if) + +lemma bind_case_let_cong [fundef_cong]: + assumes "M = N" + and "(\x. x = N \ f x s = g x s)" + shows "(Let M f) s = (Let N g) s" + by (simp add: assms(1) assms(2)) + +lemma bind_case_some_cong [fundef_cong]: + assumes "x = x'" and "\a. x = Some a \ f a s = f' a s" and "x = None \ g s = g' s" + shows "(case x of Some a \ f a | None \ g) s = (case x' of Some a \ f' a | None \ g') s" + by (simp add: assms(1) assms(2) assms(3) option.case_eq_if) + +lemma bind_case_bool_cong [fundef_cong]: + assumes "x = x'" and "x = True \ f s = f' s" and "x = False \ g s = g' s" + shows "(case x of True \ f | False \ g) s = (case x' of True \ f' | False \ g') s" + using assms(1) assms(2) assms(3) by auto + +subsection \Other functions\ + +text \ + The basic accessor functions of the state monad. \get\ returns + the current state as result, does not fail, and does not change the state. + \put s\ returns unit, changes the current state to \s\ and does not fail. +\ +fun get :: "('s, 'e, 's) state_monad" where + "get s = Normal (s, s)" + +fun put :: "'s \ (unit, 'e, 's) state_monad" where + "put s _ = Normal ((), s)" + +text \Apply a function to the current state and return the result +without changing the state.\ +fun + applyf :: "('s \ 'a) \ ('a, 'e, 's) state_monad" where + "applyf f = get \ (\s. return (f s))" + +text \Modify the current state using the function passed in.\ +fun + modify :: "('s \ 's) \ (unit, 'e, 's) state_monad" +where "modify f = get \ (\s::'s. put (f s))" + +fun + assert :: "'e \ ('s \ bool) \ (unit, 'e, 's) state_monad" where + "assert x t = (\s. if (t s) then return () s else throw x s)" + +fun + option :: "'e \ ('s \ 'a option) \ ('a, 'e, 's) state_monad" where + "option x f = (\s. (case f s of + Some y \ return y s + | None \ throw x s))" + +subsection \Some basic examples\ + +lemma "do { + x \ return 1; + return (2::nat); + return x + } = + return 1 \ (\x. return (2::nat) \ (\_. (return x)))" .. + +lemma "do { + x \ return 1; + return 2; + return x + } = return 1" + by auto + +fun sub1 :: "(unit, nat, nat) state_monad" where + "sub1 0 = put 0 0" + | "sub1 (Suc n) = (do { + x \ get; + put x; + sub1 + }) n" + +fun sub2 :: "(unit, nat, nat) state_monad" where + "sub2 s = + (do { + n \ get; + (case n of + 0 \ put 0 + | Suc n' \ (do { + put n'; + sub2 + })) + }) s" + +section "Hoare Logic" + +named_theorems wprule + +definition + valid :: "('s \ bool) \ ('a,'e,'s) state_monad \ + ('a \ 's \ bool) \ + ('e \ bool) \ bool" + ("\_\/ _ /(\_\,/ \_\)") +where + "\P\ f \Q\,\E\ \ \s. P s \ (case f s of + Normal (r,s') \ Q r s' + | Exception e \ E e)" + +lemma weaken: + assumes "\Q\ f \R\, \E\" + and "\s. P s \ Q s" + shows "\P\ f \R\, \E\" + using assms by (simp add: valid_def) + +lemma strengthen: + assumes "\P\ f \Q\, \E\" + and "\a s. Q a s \ R a s" + shows "\P\ f \R\, \E\" + unfolding valid_def +proof(rule allI[OF impI]) + fix s + assume "P s" + show "case f s of Normal (r, s') \ R r s' | Exception e \ E e" + proof (cases "f s") + case (n a s') + then show ?thesis using assms valid_def \P s\ + by (metis case_prod_conv result.simps(5)) + next + case (e e) + then show ?thesis using assms valid_def \P s\ + by (metis result.simps(6)) + qed +qed + +definition wp + where "wp f P E s \ (case f s of + Normal (r,s') \ P r s' + | Exception e \ E e)" + +lemma wp_valid: assumes "\s. P s \ (wp f Q E s)" shows "\P\ f \Q\,\E\" + unfolding valid_def by (metis assms wp_def) + +lemma valid_wp: assumes "\P\ f \Q\,\E\" shows "\s. P s \ (wp f Q E s)" + by (metis assms valid_def wp_def) + +lemma put: "\\s. P () x\ put x \P\,\E\" + using valid_def by fastforce + +lemma put': + assumes "\s. P s \ Q () x" + shows "\\s. P s\ put x \Q\,\E\" + using assms weaken[OF put, of P Q x E] by blast + +lemma wpput[wprule]: + assumes "P () x" + shows "wp (put x) P E s" + unfolding wp_def using assms by simp + +lemma get: "\\s. P s s\ get \P\,\E\" + using valid_def by fastforce + +lemma get': + assumes "\s. P s \ Q s s" + shows "\\s. P s\ get \Q\,\E\" + using assms weaken[OF get] by blast + +lemma wpget[wprule]: + assumes "P s s" + shows "wp get P E s" + unfolding wp_def using assms by simp + +lemma return: "\\s. P x s\ return x \P\,\E\" + using valid_def by fastforce + +lemma return': + assumes "\s. P s \ Q x s" + shows "\\s. P s\ return x \Q\,\E\" + using assms weaken[OF return, of P Q x E] by blast + +lemma wpreturn[wprule]: + assumes "P x s" + shows "wp (return x) P E s" + unfolding wp_def using assms by simp + +lemma bind: + assumes "\x. \B x\ g x \C\,\E\" + and "\A\ f \B\,\E\" + shows "\A\ f \ g \C\,\E\" + unfolding valid_def +proof (rule allI[OF impI]) + fix s assume a: "A s" + show "case (f \ g) s of Normal (r, s') \ C r s' | Exception e \ E e" + proof (cases "f s" rule:result_cases) + case nf: (n a s') + with assms(2) a have b: "B a s'" using valid_def[where ?f=f] by auto + then show ?thesis + proof (cases "g a s'" rule:result_cases) + case ng: (n a' s'') + with assms(1) b have c: "C a' s''" using valid_def[where ?f="g a"] by fastforce + moreover from ng nf have "(f \ g) s = Normal (a', s'')" by simp + ultimately show ?thesis by simp + next + case eg: (e e) + with assms(1) b have c: "E e" using valid_def[where ?f="g a"] by fastforce + moreover from eg nf have "(f \ g) s = Exception e" by simp + ultimately show ?thesis by simp + qed + next + case (e e) + with a assms(2) have "E e" using valid_def[where ?f=f] by auto + moreover from e have "(f \ g) s = Exception e" by simp + ultimately show ?thesis by simp + qed +qed + +lemma wpbind[wprule]: + assumes "wp f (\a. (wp (g a) P E)) E s" + shows "wp (f \ g) P E s" +proof (cases "f s" rule:result_cases) + case nf: (n a s') + then have **:"wp (g a) P E s'" using wp_def[of f "\a. wp (g a) P E"] assms by simp + show ?thesis + proof (cases "g a s'" rule:result_cases) + case ng: (n a' s'') + then have "P a' s''" using wp_def[of "g a" P] ** by simp + moreover from nf ng have "(f \ g) s = Normal (a', s'')" by simp + ultimately show ?thesis using wp_def by fastforce + next + case (e e) + then have "E e" using wp_def[of "g a" P] ** by simp + moreover from nf e have "(f \ g) s = Exception e" by simp + ultimately show ?thesis using wp_def by fastforce + qed +next + case (e e) + then have "E e" using wp_def[of f "\a. wp (g a) P E"] assms by simp + moreover from e have "(f \ g) s = Exception e" by simp + ultimately show ?thesis using wp_def by fastforce +qed + +lemma wpassert[wprule]: + assumes "t s \ wp (return ()) P E s" + and "\ t s \ wp (throw x) P E s" + shows "wp (assert x t) P E s" + using assms unfolding wp_def by simp + +lemma throw: + assumes "E x" + shows "\P\ throw x \Q\, \E\" + using valid_def assms by fastforce + +lemma wpthrow[wprule]: + assumes "E x" + shows "wp (throw x) P E s" + unfolding wp_def using assms by simp + +lemma applyf: + "\\s. P (f s) s\ applyf f \\a s. P a s\,\E\" + by (simp add: valid_def) + +lemma applyf': + assumes "\s. P s \ Q (f s) s" + shows "\\s. P s\ applyf f \\a s. Q a s\,\E\" + using assms weaken[OF applyf] by blast + +lemma wpapplyf[wprule]: + assumes "P (f s) s" + shows "wp (applyf f) P E s" + unfolding wp_def using assms by simp + +lemma modify: + "\\s. P () (f s)\ modify f \P\, \E\" + apply simp + apply (rule bind,rule allI) + apply (rule put) + apply (rule get) + done + +lemma modify': + assumes "\s. P s \ Q () (f s)" + shows "\\s. P s\ modify f \Q\, \E\" + using assms weaken[OF modify, of P Q _ E] by blast + +lemma wpmodify[wprule]: + assumes "P () (f s)" + shows "wp (modify f) P E s" + unfolding wp_def using assms by simp + +lemma wpcasenat[wprule]: + assumes "(y=(0::nat) \ wp (f y) P E s)" + and "\x. y=Suc x \ wp (g x) P E s" + shows "wp (case y::nat of 0 \ f y | Suc x \ g x) P E s" + by (metis assms(1) assms(2) not0_implies_Suc old.nat.simps(4) old.nat.simps(5)) + +lemma wpif[wprule]: + assumes "c \ wp f P E s" + and "\c \ wp g P E s" + shows "wp (if c then f else g) P E s" + using assms by simp + +lemma wpsome[wprule]: + assumes "\y. x = Some y \ wp (f y) P E s" + and "x = None \ wp g P E s" + shows "wp (case x of Some y \ f y | None \ g) P E s" + using assms unfolding wp_def by (simp split: option.split) + +lemma wpoption[wprule]: + assumes "\y. f s = Some y \ wp (return y) P E s" + and "f s = None \ wp (throw x) P E s" + shows "wp (option x f) P E s" + using assms unfolding wp_def by (auto split:option.split) + +lemma wpprod[wprule]: + assumes "\x y. a = (x,y) \ wp (f x y) P E s" + shows "wp (case a of (x, y) \ f x y) P E s" + using assms unfolding wp_def + by (simp split: prod.split) + +method wp = rule wprule; wp? +method wpvcg = rule wp_valid, wp + +lemma "\\s. s=5\ do { + put (5::nat); + x \ get; + return x + } \\a s. s=5\,\\e. False\" + by (wpvcg, simp) +end diff --git a/thys/Solidity/Statements.thy b/thys/Solidity/Statements.thy --- a/thys/Solidity/Statements.thy +++ b/thys/Solidity/Statements.thy @@ -1,3927 +1,2651 @@ -section \Statements\ -theory Statements - imports Environment StateMonad -begin - -subsection \Syntax\ - -subsubsection \Expressions\ -datatype L = Id Identifier - | Ref Identifier "E list" -and E = INT nat int - | UINT nat int - | ADDRESS String.literal - | BALANCE E - | THIS - | SENDER - | VALUE - | TRUE - | FALSE - | LVAL L - | PLUS E E - | MINUS E E - | EQUAL E E - | LESS E E - | AND E E - | OR E E - | NOT E - | CALL Identifier "E list" - | ECALL E Identifier "E list" E - -subsubsection \Statements\ -datatype S = SKIP - | BLOCK "(Identifier \ Type) \ (E option)" S - | ASSIGN L E - | TRANSFER E E - | COMP S S - | ITE E S S - | WHILE E S - | INVOKE Identifier "E list" - | EXTERNAL E Identifier "E list" E - -abbreviation - "vbits\{8,16,24,32,40,48,56,64,72,80,88,96,104,112,120,128, - 136,144,152,160,168,176,184,192,200,208,216,224,232,240,248,256}" - -lemma vbits_max[simp]: - assumes "b1 \ vbits" - and "b2 \ vbits" - shows "(max b1 b2) \ vbits" -proof - - consider (b1) "max b1 b2 = b1" | (b2) "max b1 b2 = b2" by (metis max_def) - then show ?thesis - proof cases - case b1 - then show ?thesis using assms(1) by simp - next - case b2 - then show ?thesis using assms(2) by simp - qed -qed - -lemma vbits_ge_0: "(x::nat)\vbits \ x>0" by auto - -subsection \Contracts\ - -text \ - A contract consists of methods or storage variables. - A method is a triple consisting of - \begin{itemize} - \item A list of formal parameters - \item A statement - \item An optional return value - \end{itemize} -\ - -datatype Member = Method "(Identifier \ Type) list \ S \ E option" - | Var STypes - -text \ - A procedure environment assigns a contract to an address. - A contract consists of - \begin{itemize} - \item An assignment of members to identifiers - \item An optional fallback statement which is executed after money is beeing transferred to the contract. - \end{itemize} - \url{https://docs.soliditylang.org/en/v0.8.6/contracts.html#fallback-function} -\ - -type_synonym Environment\<^sub>P = "(Address, (Identifier, Member) fmap \ S) fmap" - -definition init::"(Identifier, Member) fmap \ Identifier \ Environment \ Environment" - where "init ct i e = (case fmlookup ct i of - Some (Var tp) \ updateEnvDup i (Storage tp) (Storeloc i) e - | _ \ e)" - -lemma init_s11[simp]: - assumes "fmlookup ct i = Some (Var tp)" - shows "init ct i e = updateEnvDup i (Storage tp) (Storeloc i) e" - using assms init_def by simp - -lemma init_s12[simp]: - assumes "i |\| fmdom (denvalue e)" - shows "init ct i e = e" -proof (cases "fmlookup ct i") - case None - then show ?thesis using init_def by simp -next - case (Some a) - then show ?thesis - proof (cases a) - case (Method x1) - with Some show ?thesis using init_def by simp - next - case (Var tp) - with Some have "init ct i e = updateEnvDup i (Storage tp) (Storeloc i) e" using init_def by simp - moreover from assms have "updateEnvDup i (Storage tp) (Storeloc i) e = e" by auto - ultimately show ?thesis by simp - qed -qed - -lemma init_s13[simp]: - assumes "fmlookup ct i = Some (Var tp)" - and "\ i |\| fmdom (denvalue e)" - shows "init ct i e = updateEnv i (Storage tp) (Storeloc i) e" - using assms init_def by auto - -lemma init_s21[simp]: - assumes "fmlookup ct i = None" - shows "init ct i e = e" - using assms init_def by auto - -lemma init_s22[simp]: - assumes "fmlookup ct i = Some (Method m)" - shows "init ct i e = e" - using assms init_def by auto - -lemma init_commte: "comp_fun_commute (init ct)" -proof - fix x y - show "init ct y \ init ct x = init ct x \ init ct y" - proof - fix e - show "(init ct y \ init ct x) e = (init ct x \ init ct y) e" - proof (cases "fmlookup ct x") - case None - then show ?thesis by simp - next - case s1: (Some a) - then show ?thesis - proof (cases a) - case (Method x1) - with s1 show ?thesis by simp - next - case v1: (Var tp) - then show ?thesis - proof (cases "x |\| fmdom (denvalue e)") - case True - with s1 v1 have *: "init ct x e = e" by auto - then show ?thesis - proof (cases "fmlookup ct y") - case None - then show ?thesis by simp - next - case s2: (Some a) - then show ?thesis - proof (cases a) - case (Method x1) - with s2 show ?thesis by simp - next - case v2: (Var tp') - then show ?thesis - proof (cases "y |\| fmdom (denvalue e)") - case t1: True - with s1 v1 True s2 v2 show ?thesis by fastforce - next - define e' where "e' = updateEnv y (Storage tp') (Storeloc y) e" - case False - with s2 v2 have "init ct y e = e'" using e'_def by auto - with s1 v1 True e'_def * show ?thesis by auto - qed - qed - qed - next - define e' where "e' = updateEnv x (Storage tp) (Storeloc x) e" - case f1: False - with s1 v1 have *: "init ct x e = e'" using e'_def by auto - then show ?thesis - proof (cases "fmlookup ct y") - case None - then show ?thesis by simp - next - case s3: (Some a) - then show ?thesis - proof (cases a) - case (Method x1) - with s3 show ?thesis by simp - next - case v2: (Var tp') - then show ?thesis - proof (cases "y |\| fmdom (denvalue e)") - case t1: True - with e'_def have "y |\| fmdom (denvalue e')" by simp - with s1 s3 v1 f1 v2 show ?thesis using e'_def by fastforce - next - define f' where "f' = updateEnv y (Storage tp') (Storeloc y) e" - define e'' where "e'' = updateEnv y (Storage tp') (Storeloc y) e'" - case f2: False - with s3 v2 have **: "init ct y e = f'" using f'_def by auto - show ?thesis - proof (cases "y = x") - case True - with s3 v2 e'_def have "init ct y e' = e'" by simp - moreover from s3 v2 True f'_def have "init ct x f' = f'" by simp - ultimately show ?thesis using True by simp - next - define f'' where "f'' = updateEnv x (Storage tp) (Storeloc x) f'" - case f3: False - with f2 have "\ y |\| fmdom (denvalue e')" using e'_def by simp - with s3 v2 e''_def have "init ct y e' = e''" by auto - with * have "(init ct y \ init ct x) e = e''" by simp - moreover have "init ct x f' = f''" - proof - - from s1 v1 have "init ct x f' = updateEnvDup x (Storage tp) (Storeloc x) f'" by simp - moreover from f1 f3 have "x |\| fmdom (denvalue f')" using f'_def by simp - ultimately show ?thesis using f''_def by auto - qed - moreover from f''_def e''_def f'_def e'_def f3 have "Some f'' = Some e''" using env_reorder_neq by simp - ultimately show ?thesis using ** by simp - qed - qed - qed - qed - qed - qed - qed - qed -qed - -lemma init_address[simp]: - "address (init ct i e) = address e \ sender (init ct i e) = sender e" -proof (cases "fmlookup ct i") - case None - then show ?thesis by simp -next - case (Some a) - show ?thesis - proof (cases a) - case (Method x1) - with Some show ?thesis by simp - next - case (Var tp) - with Some show ?thesis using updateEnvDup_address updateEnvDup_sender by simp - qed -qed - -lemma init_sender[simp]: -"sender (init ct i e) = sender e" -proof (cases "fmlookup ct i") - case None - then show ?thesis by simp -next - case (Some a) - show ?thesis - proof (cases a) - case (Method x1) - with Some show ?thesis by simp - next - case (Var tp) - with Some show ?thesis using updateEnvDup_sender by simp - qed -qed - -lemma init_svalue[simp]: -"svalue (init ct i e) = svalue e" -proof (cases "fmlookup ct i") - case None - then show ?thesis by simp -next - case (Some a) - show ?thesis - proof (cases a) - case (Method x1) - with Some show ?thesis by simp - next - case (Var tp) - with Some show ?thesis using updateEnvDup_svalue by simp - qed -qed - -lemma ffold_init_ad_same[rule_format]: "\e'. ffold (init ct) e xs = e' \ address e' = address e \ sender e' = sender e \ svalue e' = svalue e" -proof (induct xs) - case empty - then show ?case by (simp add: ffold_def) -next - case (insert x xs) - then have *: "ffold (init ct) e (finsert x xs) = - init ct x (ffold (init ct) e xs)" using FSet.comp_fun_commute.ffold_finsert[OF init_commte] by simp - show ?case - proof (rule allI[OF impI]) - fix e' assume **: "ffold (init ct) e (finsert x xs) = e'" - with * obtain e'' where ***: "ffold (init ct) e xs = e''" by simp - with insert have "address e'' = address e \ sender e'' = sender e \ svalue e'' = svalue e" by blast - with * ** *** show "address e' = address e \ sender e' = sender e \ svalue e' = svalue e" using init_address init_sender init_svalue by metis - qed -qed - -lemma ffold_init_dom: - "fmdom (denvalue (ffold (init ct) e xs)) |\| fmdom (denvalue e) |\| xs" -proof (induct "xs") - case empty - then show ?case - proof - fix x - assume "x |\| fmdom (denvalue (ffold (init ct) e {||}))" - moreover have "ffold (init ct) e {||} = e" - using FSet.comp_fun_commute.ffold_empty[OF init_commte, of ct e] by simp - ultimately show "x |\| fmdom (denvalue e) |\| {||}" by simp - qed -next - case (insert x xs) - then have *: "ffold (init ct) e (finsert x xs) = - init ct x (ffold (init ct) e xs)" using FSet.comp_fun_commute.ffold_finsert[OF init_commte] by simp - - show ?case - proof - fix x' assume "x' |\| fmdom (denvalue (ffold (init ct) e (finsert x xs)))" - with * have **: "x' |\| fmdom (denvalue (init ct x (ffold (init ct) e xs)))" by simp - then consider "x' |\| fmdom (denvalue (ffold (init ct) e xs))" | "x'=x" - proof (cases "fmlookup ct x") - case None - then show ?thesis using that ** by simp - next - case (Some a) - then show ?thesis - proof (cases a) - case (Method x1) - then show ?thesis using Some ** that by simp - next - case (Var x2) - show ?thesis - proof (cases "x=x'") - case True - then show ?thesis using that by simp - next - case False - then have "fmlookup (denvalue (updateEnvDup x (Storage x2) (Storeloc x) (ffold (init ct) e xs))) x' = fmlookup (denvalue (ffold (init ct) e xs)) x'" using updateEnvDup_dup by simp - moreover from ** Some Var have ***:"x' |\| fmdom (denvalue (updateEnvDup x (Storage x2) (Storeloc x) (ffold (init ct) e xs)))" by simp - ultimately have "x' |\| fmdom (denvalue (ffold (init ct) e xs))" by (simp add: fmlookup_dom_iff) - then show ?thesis using that by simp - qed - qed - qed - then show "x' |\| fmdom (denvalue e) |\| finsert x xs" - proof cases - case 1 - then show ?thesis using insert.hyps by auto - next - case 2 - then show ?thesis by simp - qed - qed -qed - -lemma ffold_init_fmap: - assumes "fmlookup ct i = Some (Var tp)" - and "i |\| fmdom (denvalue e)" - shows "i|\|xs \ fmlookup (denvalue (ffold (init ct) e xs)) i = Some (Storage tp, Storeloc i)" -proof (induct "xs") - case empty - then show ?case by simp -next - case (insert x xs) - then have *: "ffold (init ct) e (finsert x xs) = - init ct x (ffold (init ct) e xs)" using FSet.comp_fun_commute.ffold_finsert[OF init_commte] by simp - - from insert.prems consider (a) "i |\| xs" | (b) "\ i |\| xs \ i = x" by auto - then show "fmlookup (denvalue (ffold (init ct) e (finsert x xs))) i = Some (Storage tp, Storeloc i)" - proof cases - case a - with insert.hyps(2) have "fmlookup (denvalue (ffold (init ct) e xs)) i = Some (Storage tp, Storeloc i)" by simp - moreover have "fmlookup (denvalue (init ct x (ffold (init ct) e xs))) i = fmlookup (denvalue (ffold (init ct) e xs)) i" - proof (cases "fmlookup ct x") - case None - then show ?thesis by simp - next - case (Some a) - then show ?thesis - proof (cases a) - case (Method x1) - with Some show ?thesis by simp - next - case (Var x2) - with Some have "init ct x (ffold (init ct) e xs) = updateEnvDup x (Storage x2) (Storeloc x) (ffold (init ct) e xs)" using init_def[of ct x "(ffold (init ct) e xs)"] by simp - moreover from insert a have "i\x" by auto - then have "fmlookup (denvalue (updateEnvDup x (Storage x2) (Storeloc x) (ffold (init ct) e xs))) i = fmlookup (denvalue (ffold (init ct) e xs)) i" using updateEnvDup_dup[of x i] by simp - ultimately show ?thesis by simp - qed - qed - ultimately show ?thesis using * by simp - next - case b - with assms(1) have "fmlookup ct x = Some (Var tp)" by simp - moreover from b assms(2) have "\ x |\| fmdom (denvalue (ffold (init ct) e xs))" using ffold_init_dom by auto - ultimately have "init ct x (ffold (init ct) e xs) = updateEnv x (Storage tp) (Storeloc x) (ffold (init ct) e xs)" by auto - with b * show ?thesis by simp - qed -qed - -text\The following definition allows for a more fine-grained configuration of the - code generator. -\ -definition ffold_init::"(String.literal, Member) fmap \ Environment \ String.literal fset \ Environment" where - \ffold_init ct a c = ffold (init ct) a c\ -declare ffold_init_def [simp] - -lemma ffold_init_code [code]: - \ffold_init ct a c = fold (init ct) (remdups (sorted_list_of_set (fset c))) a\ - using comp_fun_commute_on.fold_set_fold_remdups ffold.rep_eq - ffold_init_def init_commte sorted_list_of_fset.rep_eq - sorted_list_of_fset_simps(1) - by (metis comp_fun_commute.comp_fun_commute comp_fun_commute_on.intro order_refl) - -lemma bind_case_stackvalue_cong [fundef_cong]: - assumes "x = x'" - and "\v. x = KValue v \ f v s = f' v s" - and "\p. x = KCDptr p \ g p s = g' p s" - and "\p. x = KMemptr p \ h p s = h' p s" - and "\p. x = KStoptr p \ i p s = i' p s" - shows "(case x of KValue v \ f v | KCDptr p \ g p | KMemptr p \ h p | KStoptr p \ i p) s - = (case x' of KValue v \ f' v | KCDptr p \ g' p | KMemptr p \ h' p | KStoptr p \ i' p) s" - using assms by (cases x, auto) - -lemma bind_case_type_cong [fundef_cong]: - assumes "x = x'" - and "\t. x = Value t \ f t s = f' t s" - and "\t. x = Calldata t \ g t s = g' t s" - and "\t. x = Memory t \ h t s = h' t s" - and "\t. x = Storage t \ i t s = i' t s" - shows "(case x of Value t \ f t | Calldata t \ g t | Memory t \ h t | Storage t \ i t ) s - = (case x' of Value t \ f' t | Calldata t \ g' t | Memory t \ h' t | Storage t \ i' t) s" - using assms by (cases x, auto) - -lemma bind_case_denvalue_cong [fundef_cong]: - assumes "x = x'" - and "\a. x = (Stackloc a) \ f a s = f' a s" - and "\a. x = (Storeloc a) \ g a s = g' a s" - shows "(case x of (Stackloc a) \ f a | (Storeloc a) \ g a) s - = (case x' of (Stackloc a) \ f' a | (Storeloc a) \ g' a) s" - using assms by (cases x, auto) - -lemma bind_case_mtypes_cong [fundef_cong]: - assumes "x = x'" - and "\a t. x = (MTArray a t) \ f a t s = f' a t s" - and "\p. x = (MTValue p) \ g p s = g' p s" - shows "(case x of (MTArray a t) \ f a t | (MTValue p) \ g p) s - = (case x' of (MTArray a t) \ f' a t | (MTValue p) \ g' p) s" - using assms by (cases x, auto) - -lemma bind_case_stypes_cong [fundef_cong]: - assumes "x = x'" - and "\a t. x = (STArray a t) \ f a t s = f' a t s" - and "\a t. x = (STMap a t) \ g a t s = g' a t s" - and "\p. x = (STValue p) \ h p s = h' p s" - shows "(case x of (STArray a t) \ f a t | (STMap a t) \ g a t | (STValue p) \ h p) s - = (case x' of (STArray a t) \ f' a t | (STMap a t) \ g' a t | (STValue p) \ h' p) s" - using assms by (cases x, auto) - -lemma bind_case_types_cong [fundef_cong]: - assumes "x = x'" - and "\a. x = (TSInt a) \ f a s = f' a s" - and "\a. x = (TUInt a) \ g a s = g' a s" - and "x = TBool \ h s = h' s" - and "x = TAddr \ i s = i' s" - shows "(case x of (TSInt a) \ f a | (TUInt a) \ g a | TBool \ h | TAddr \ i) s - = (case x' of (TSInt a) \ f' a | (TUInt a) \ g' a | TBool \ h' | TAddr \ i') s" - using assms by (cases x, auto) - -lemma bind_case_contract_cong [fundef_cong]: - assumes "x = x'" - and "\a. x = Method a \ f a s = f' a s" - and "\a. x = Var a \ g a s = g' a s" - shows "(case x of (Method a) \ f a | (Var a) \ g a) s - = (case x' of (Method a) \ f' a | (Var a) \ g' a) s" - using assms by (cases x, auto) - -lemma bind_case_memoryvalue_cong [fundef_cong]: - assumes "x = x'" - and "\a. x = MValue a \ f a s = f' a s" - and "\a. x = MPointer a \ g a s = g' a s" - shows "(case x of (MValue a) \ f a | (MPointer a) \ g a) s - = (case x' of (MValue a) \ f' a | (MPointer a) \ g' a) s" - using assms by (cases x, auto) - -abbreviation lift :: - "(E \ Environment\<^sub>P \ Environment \ CalldataT \ (Stackvalue * Type, Ex, State) state_monad) - \ (Types \ Types \ Valuetype \ Valuetype \ (Valuetype * Types) option) - \ E \ E \ Environment\<^sub>P \ Environment \ CalldataT \ (Stackvalue * Type, Ex, State) state_monad" -where - "lift expr f e1 e2 e\<^sub>p e cd \ - (do { - kv1 \ expr e1 e\<^sub>p e cd; - (case kv1 of - (KValue v1, Value t1) \ (do - { - kv2 \ expr e2 e\<^sub>p e cd; - (case kv2 of - (KValue v2, Value t2) \ - (case f t1 t2 v1 v2 of - Some (v, t) \ return (KValue v, Value t) - | None \ throw Err) - | _ \ (throw Err::(Stackvalue * Type, Ex, State) state_monad)) - }) - | _ \ (throw Err::(Stackvalue * Type, Ex, State) state_monad)) - })" - -abbreviation gascheck :: - "(State \ Gas) \ (unit, Ex, State) state_monad" -where - "gascheck check \ - do { - g \ (applyf check::(Gas, Ex, State) state_monad); - (assert Gas (\st. gas st \ g) (modify (\st. st \gas:=gas st - g\)::(unit, Ex, State) state_monad)) - }" - -subsection \Semantics\ - -datatype LType = LStackloc Location - | LMemloc Location - | LStoreloc Location - - -locale statement_with_gas = - fixes costs :: "S\ Environment\<^sub>P \ Environment \ CalldataT \ State \ Gas" - and costs\<^sub>e :: "E\ Environment\<^sub>P \ Environment \ CalldataT \ State \ Gas" - assumes while_not_zero[termination_simp]: "\e e\<^sub>p cd st ex s0. 0 < (costs (WHILE ex s0) e\<^sub>p e cd st) " - and call_not_zero[termination_simp]: "\e e\<^sub>p cd st i ix. 0 < (costs\<^sub>e (CALL i ix) e\<^sub>p e cd st)" - and ecall_not_zero[termination_simp]: "\e e\<^sub>p cd st a i ix val. 0 < (costs\<^sub>e (ECALL a i ix val) e\<^sub>p e cd st)" - and invoke_not_zero[termination_simp]: "\e e\<^sub>p cd st i xe. 0 < (costs (INVOKE i xe) e\<^sub>p e cd st)" - and external_not_zero[termination_simp]: "\e e\<^sub>p cd st ad i xe val. 0 < (costs (EXTERNAL ad i xe val) e\<^sub>p e cd st)" - and transfer_not_zero[termination_simp]: "\e e\<^sub>p cd st ex ad. 0 < (costs (TRANSFER ad ex) e\<^sub>p e cd st)" -begin - -function msel::"bool \ MTypes \ Location \ E list \ Environment\<^sub>P \ Environment \ CalldataT \ (Location * MTypes, Ex, State) state_monad" - and ssel::"STypes \ Location \ E list \ Environment\<^sub>P \ Environment \ CalldataT \ (Location * STypes, Ex, State) state_monad" - and lexp :: "L \ Environment\<^sub>P \ Environment \ CalldataT \ (LType * Type, Ex, State) state_monad" - and expr::"E \ Environment\<^sub>P \ Environment \ CalldataT \ (Stackvalue * Type, Ex, State) state_monad" - and load :: "bool \ (Identifier \ Type) list \ E list \ Environment\<^sub>P \ Environment \ CalldataT \ State \ Environment \ CalldataT \ (Environment \ CalldataT \ State, Ex, State) state_monad" - and rexp::"L \ Environment\<^sub>P \ Environment \ CalldataT \ (Stackvalue * Type, Ex, State) state_monad" - and stmt :: "S \ Environment\<^sub>P \ Environment \ CalldataT \ (unit, Ex, State) state_monad" -where - "msel _ _ _ [] _ _ _ st = throw Err st" -| "msel _ (MTValue _) _ _ _ _ _ st = throw Err st" -| "msel _ (MTArray al t) loc [x] e\<^sub>p env cd st = - (do { - kv \ expr x e\<^sub>p env cd; - (case kv of - (KValue v, Value t') \ - (if less t' (TUInt 256) v (ShowL\<^sub>i\<^sub>n\<^sub>t al) = Some (ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True, TBool) - then return (hash loc v, t) - else throw Err) - | _ \ throw Err) - }) st" -(* - Note that it is indeed possible to modify the state while evaluating the expression - to determine the index of the array to look up. -*) -| "msel mm (MTArray al t) loc (x # y # ys) e\<^sub>p env cd st = - (do { - kv \ expr x e\<^sub>p env cd; - (case kv of - (KValue v, Value t') \ - (if less t' (TUInt 256) v (ShowL\<^sub>i\<^sub>n\<^sub>t al) = Some (ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True, TBool) - then do { - s \ applyf (\st. if mm then memory st else cd); - (case accessStore (hash loc v) s of - Some (MPointer l) \ msel mm t l (y#ys) e\<^sub>p env cd - | _ \ throw Err) - } else throw Err) - | _ \ throw Err) - }) st" -| "ssel tp loc Nil _ _ _ st = return (loc, tp) st" -| "ssel (STValue _) _ (_ # _) _ _ _ st = throw Err st" -| "ssel (STArray al t) loc (x # xs) e\<^sub>p env cd st = - (do { - kv \ expr x e\<^sub>p env cd; - (case kv of - (KValue v, Value t') \ - (if less t' (TUInt 256) v (ShowL\<^sub>i\<^sub>n\<^sub>t al) = Some (ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True, TBool) - then ssel t (hash loc v) xs e\<^sub>p env cd - else throw Err) - | _ \ throw Err) - }) st" -| "ssel (STMap _ t) loc (x # xs) e\<^sub>p env cd st = - (do { - kv \ expr x e\<^sub>p env cd; - (case kv of - (KValue v, _) \ ssel t (hash loc v) xs e\<^sub>p env cd - | _ \ throw Err) - }) st" -| "lexp (Id i) _ e _ st = - (case fmlookup (denvalue e) i of - Some (tp, (Stackloc l)) \ return (LStackloc l, tp) - | Some (tp, (Storeloc l)) \ return (LStoreloc l, tp) - | _ \ throw Err) st" -| "lexp (Ref i r) e\<^sub>p e cd st = - (case fmlookup (denvalue e) i of - Some (tp, Stackloc l) \ - do { - k \ applyf (\st. accessStore l (stack st)); - (case k of - Some (KCDptr _) \ throw Err - | Some (KMemptr l') \ - (case tp of - Memory t \ - do { - (l'', t') \ msel True t l' r e\<^sub>p e cd; - return (LMemloc l'', Memory t') - } - | _ \ throw Err) - | Some (KStoptr l') \ - (case tp of - Storage t \ - do { - (l'', t') \ ssel t l' r e\<^sub>p e cd; - return (LStoreloc l'', Storage t') - } - | _ \ throw Err) - | Some (KValue _) \ throw Err - | None \ throw Err) - } - | Some (tp, Storeloc l) \ - (case tp of - Storage t \ - do { - (l', t') \ ssel t l r e\<^sub>p e cd; - return (LStoreloc l', Storage t') - } - | _ \ throw Err) - | None \ throw Err) st" -| "expr (E.INT b x) e\<^sub>p e cd st = - (do { - gascheck (costs\<^sub>e (E.INT b x) e\<^sub>p e cd); - (if (b \ vbits) - then (return (KValue (createSInt b x), Value (TSInt b))) - else (throw Err)) - }) st" -| "expr (UINT b x) e\<^sub>p e cd st = - (do { - gascheck (costs\<^sub>e (UINT b x) e\<^sub>p e cd); - (if (b \ vbits) - then (return (KValue (createUInt b x), Value (TUInt b))) - else (throw Err)) - }) st" -| "expr (ADDRESS ad) e\<^sub>p e cd st = - (do { - gascheck (costs\<^sub>e (ADDRESS ad) e\<^sub>p e cd); - return (KValue ad, Value TAddr) - }) st" -| "expr (BALANCE ad) e\<^sub>p e cd st = - (do { - gascheck (costs\<^sub>e (BALANCE ad) e\<^sub>p e cd); - kv \ expr ad e\<^sub>p e cd; - (case kv of - (KValue adv, Value TAddr) \ - return (KValue (accessBalance (accounts st) adv), Value (TUInt 256)) - | _ \ throw Err) - }) st" -| "expr THIS e\<^sub>p e cd st = - (do { - gascheck (costs\<^sub>e THIS e\<^sub>p e cd); - return (KValue (address e), Value TAddr) - }) st" -| "expr SENDER e\<^sub>p e cd st = - (do { - gascheck (costs\<^sub>e SENDER e\<^sub>p e cd); - return (KValue (sender e), Value TAddr) - }) st" -| "expr VALUE e\<^sub>p e cd st = - (do { - gascheck (costs\<^sub>e VALUE e\<^sub>p e cd); - return (KValue (svalue e), Value (TUInt 256)) - }) st" -| "expr TRUE e\<^sub>p e cd st = - (do { - gascheck (costs\<^sub>e TRUE e\<^sub>p e cd); - return (KValue (ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True), Value TBool) - }) st" -| "expr FALSE e\<^sub>p e cd st = - (do { - gascheck (costs\<^sub>e FALSE e\<^sub>p e cd); - return (KValue (ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l False), Value TBool) - }) st" -| "expr (NOT x) e\<^sub>p e cd st = - (do { - gascheck (costs\<^sub>e (NOT x) e\<^sub>p e cd); - kv \ expr x e\<^sub>p e cd; - (case kv of - (KValue v, Value t) \ - (if v = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True - then expr FALSE e\<^sub>p e cd - else (if v = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l False - then expr TRUE e\<^sub>p e cd - else throw Err)) - | _ \ throw Err) - }) st" -| "expr (PLUS e1 e2) e\<^sub>p e cd st = (gascheck (costs\<^sub>e (PLUS e1 e2) e\<^sub>p e cd) \ (\_. lift expr add e1 e2 e\<^sub>p e cd)) st" -| "expr (MINUS e1 e2) e\<^sub>p e cd st = (gascheck (costs\<^sub>e (MINUS e1 e2) e\<^sub>p e cd) \ (\_. lift expr sub e1 e2 e\<^sub>p e cd)) st" -| "expr (LESS e1 e2) e\<^sub>p e cd st = (gascheck (costs\<^sub>e (LESS e1 e2) e\<^sub>p e cd) \ (\_. lift expr less e1 e2 e\<^sub>p e cd)) st" -| "expr (EQUAL e1 e2) e\<^sub>p e cd st = (gascheck (costs\<^sub>e (EQUAL e1 e2) e\<^sub>p e cd) \ (\_. lift expr equal e1 e2 e\<^sub>p e cd)) st" -| "expr (AND e1 e2) e\<^sub>p e cd st = (gascheck (costs\<^sub>e (AND e1 e2) e\<^sub>p e cd) \ (\_. lift expr vtand e1 e2 e\<^sub>p e cd)) st" -| "expr (OR e1 e2) e\<^sub>p e cd st = (gascheck (costs\<^sub>e (OR e1 e2) e\<^sub>p e cd) \ (\_. lift expr vtor e1 e2 e\<^sub>p e cd)) st" -| "expr (LVAL i) e\<^sub>p e cd st = - (do { - gascheck (costs\<^sub>e (LVAL i) e\<^sub>p e cd); - rexp i e\<^sub>p e cd - }) st" -(* Notes about method calls: - - Internal method calls use a fresh environment and stack but keep the memory [1] - - External method calls use a fresh environment and stack and memory [2] - [1]: https://docs.soliditylang.org/en/v0.8.5/control-structures.html#internal-function-calls - [2]: https://docs.soliditylang.org/en/v0.8.5/control-structures.html#external-function-calls - -TODO: Functions with no return value should be able to execute -*) -| "expr (CALL i xe) e\<^sub>p e cd st = - (do { - gascheck (costs\<^sub>e (CALL i xe) e\<^sub>p e cd); - (case fmlookup e\<^sub>p (address e) of - Some (ct, _) \ - (case fmlookup ct i of - Some (Method (fp, f, Some x)) \ - let e' = ffold_init ct (emptyEnv (address e) (sender e) (svalue e)) (fmdom ct) - in (do { - st' \ applyf (\st. st\stack:=emptyStore\); - (e'', cd', st'') \ load False fp xe e\<^sub>p e' emptyStore st' e cd; - st''' \ get; - put st''; - stmt f e\<^sub>p e'' cd'; - rv \ expr x e\<^sub>p e'' cd'; - modify (\st. st\stack:=stack st''', memory := memory st'''\); - return rv - }) - | _ \ throw Err) - | None \ throw Err) - }) st" -| "expr (ECALL ad i xe val) e\<^sub>p e cd st = - (do { - gascheck (costs\<^sub>e (ECALL ad i xe val) e\<^sub>p e cd); - kad \ expr ad e\<^sub>p e cd; - (case kad of - (KValue adv, Value TAddr) \ - (case fmlookup e\<^sub>p adv of - Some (ct, _) \ - (case fmlookup ct i of - Some (Method (fp, f, Some x)) \ - (do { - kv \ expr val e\<^sub>p e cd; - (case kv of - (KValue v, Value t) \ - let e' = ffold_init ct (emptyEnv adv (address e) v) (fmdom ct) - in (do { - st' \ applyf (\st. st\stack:=emptyStore, memory:=emptyStore\); - (e'', cd', st'') \ load True fp xe e\<^sub>p e' emptyStore st' e cd; - st''' \ get; - (case transfer (address e) adv v (accounts st'') of - Some acc \ - do { - put (st''\accounts := acc\); - stmt f e\<^sub>p e'' cd'; - rv \ expr x e\<^sub>p e'' cd'; - modify (\st. st\stack:=stack st''', memory := memory st'''\); - return rv - } - | None \ throw Err) - }) - | _ \ throw Err) - }) - | _ \ throw Err) - | None \ throw Err) - | _ \ throw Err) - }) st" -| "load cp ((i\<^sub>p, t\<^sub>p)#pl) (e#el) e\<^sub>p e\<^sub>v' cd' st' e\<^sub>v cd st = - (do { - (v, t) \ expr e e\<^sub>p e\<^sub>v cd; - st'' \ get; - put st'; - (cd'', e\<^sub>v'') \ decl i\<^sub>p t\<^sub>p (Some (v,t)) cp cd (memory st'') cd' e\<^sub>v'; - st''' \ get; - put st''; - load cp pl el e\<^sub>p e\<^sub>v'' cd'' st''' e\<^sub>v cd - }) st" -| "load _ [] (_#_) _ _ _ _ _ _ st = throw Err st" -| "load _ (_#_) [] _ _ _ _ _ _ st = throw Err st" -| "load _ [] [] _ e\<^sub>v' cd' st' e\<^sub>v cd st = return (e\<^sub>v', cd', st') st" -(*TODO: Should be possible to simplify*) -| "rexp (Id i) e\<^sub>p e cd st = - (case fmlookup (denvalue e) i of - Some (tp, Stackloc l) \ - do { - s \ applyf (\st. accessStore l (stack st)); - (case s of - Some (KValue v) \ return (KValue v, tp) - | Some (KCDptr p) \ return (KCDptr p, tp) - | Some (KMemptr p) \ return (KMemptr p, tp) - | Some (KStoptr p) \ return (KStoptr p, tp) - | _ \ throw Err) - } - | Some (Storage (STValue t), Storeloc l) \ - do { - so \ applyf (\st. fmlookup (storage st) (address e)); - (case so of - Some s \ return (KValue (accessStorage t l s), Value t) - | None \ throw Err) - } - | Some (Storage (STArray x t), Storeloc l) \ return (KStoptr l, Storage (STArray x t)) - | _ \ throw Err) st" -| "rexp (Ref i r) e\<^sub>p e cd st = - (case fmlookup (denvalue e) i of - Some (tp, (Stackloc l)) \ - do { - kv \ applyf (\st. accessStore l (stack st)); - (case kv of - Some (KCDptr l') \ - (case tp of - Calldata t \ - do { - (l'', t') \ msel False t l' r e\<^sub>p e cd; - (case t' of - MTValue t'' \ - (case accessStore l'' cd of - Some (MValue v) \ return (KValue v, Value t'') - | _ \ throw Err) - | MTArray x t'' \ - (case accessStore l'' cd of - Some (MPointer p) \ return (KCDptr p, Calldata (MTArray x t'')) - | _ \ throw Err)) - } - | _ \ throw Err) - | Some (KMemptr l') \ - (case tp of - Memory t \ - do { - (l'', t') \ msel True t l' r e\<^sub>p e cd; - (case t' of - MTValue t'' \ - do { - mv \ applyf (\st. accessStore l'' (memory st)); - (case mv of - Some (MValue v) \ return (KValue v, Value t'') - | _ \ throw Err) - } - | MTArray x t'' \ - do { - mv \ applyf (\st. accessStore l'' (memory st)); - (case mv of - Some (MPointer p) \ return (KMemptr p, Memory (MTArray x t'')) - | _ \ throw Err) - } - ) - } - | _ \ throw Err) - | Some (KStoptr l') \ - (case tp of - Storage t \ - do { - (l'', t') \ ssel t l' r e\<^sub>p e cd; - (case t' of - STValue t'' \ - do { - so \ applyf (\st. fmlookup (storage st) (address e)); - (case so of - Some s \ return (KValue (accessStorage t'' l'' s), Value t'') - | None \ throw Err) - } - | STArray _ _ \ return (KStoptr l'', Storage t') - | STMap _ _ \ return (KStoptr l'', Storage t')) - } - | _ \ throw Err) - | _ \ throw Err) - } - | Some (tp, (Storeloc l)) \ - (case tp of - Storage t \ - do { - (l', t') \ ssel t l r e\<^sub>p e cd; - (case t' of - STValue t'' \ - do { - so \ applyf (\st. fmlookup (storage st) (address e)); - (case so of - Some s \ return (KValue (accessStorage t'' l' s), Value t'') - | None \ throw Err) - } - | STArray _ _ \ return (KStoptr l', Storage t') - | STMap _ _ \ return (KStoptr l', Storage t')) - } - | _ \ throw Err) - | None \ throw Err) st" -| "stmt SKIP e\<^sub>p e cd st = gascheck (costs SKIP e\<^sub>p e cd) st" -| "stmt (ASSIGN lv ex) e\<^sub>p env cd st = - (do { - gascheck (costs (ASSIGN lv ex) e\<^sub>p env cd); - re \ expr ex e\<^sub>p env cd; - (case re of - (KValue v, Value t) \ - do { - rl \ lexp lv e\<^sub>p env cd; - (case rl of - (LStackloc l, Value t') \ - (case convert t t' v of - Some (v', _) \ modify (\st. st \stack := updateStore l (KValue v') (stack st)\) - | None \ throw Err) - | (LStoreloc l, Storage (STValue t')) \ - (case convert t t' v of - Some (v', _) \ - do { - so \ applyf (\st. fmlookup (storage st) (address env)); - (case so of - Some s \ modify (\st. st\storage := fmupd (address env) (fmupd l v' s) (storage st)\) - | None \ throw Err) - } - | None \ throw Err) - | (LMemloc l, Memory (MTValue t')) \ - (case convert t t' v of - Some (v', _) \ modify (\st. st\memory := updateStore l (MValue v') (memory st)\) - | None \ throw Err) - | _ \ throw Err) - } - | (KCDptr p, Calldata (MTArray x t)) \ - do { - rl \ lexp lv e\<^sub>p env cd; - (case rl of - (LStackloc l, Memory _) \ modify (\st. st \stack := updateStore l (KCDptr p) (stack st)\) - | (LStackloc l, Storage _) \ - do { - sv \ applyf (\st. accessStore l (stack st)); - (case sv of - Some (KStoptr p') \ - do { - so \ applyf (\st. fmlookup (storage st) (address env)); - (case so of - Some s \ - (case cpm2s p p' x t cd s of - Some s' \ modify (\st. st \storage := fmupd (address env) s' (storage st)\) - | None \ throw Err) - | None \ throw Err) - } - | _ \ throw Err) - } - | (LStoreloc l, _) \ - do { - so \ applyf (\st. fmlookup (storage st) (address env)); - (case so of - Some s \ - (case cpm2s p l x t cd s of - Some s' \ modify (\st. st \storage := fmupd (address env) s' (storage st)\) - | None \ throw Err) - | None \ throw Err) - } - | (LMemloc l, _) \ - do { - cs \ applyf (\st. cpm2m p l x t cd (memory st)); - (case cs of - Some m \ modify (\st. st \memory := m\) - | None \ throw Err) - } - | _ \ throw Err) - } - | (KMemptr p, Memory (MTArray x t)) \ - do { - rl \ lexp lv e\<^sub>p env cd; - (case rl of - (LStackloc l, Memory _) \ modify (\st. st\stack := updateStore l (KMemptr p) (stack st)\) - | (LStackloc l, Storage _) \ - do { - sv \ applyf (\st. accessStore l (stack st)); - (case sv of - Some (KStoptr p') \ - do { - so \ applyf (\st. fmlookup (storage st) (address env)); - (case so of - Some s \ - do { - cs \ applyf (\st. cpm2s p p' x t (memory st) s); - (case cs of - Some s' \ modify (\st. st \storage := fmupd (address env) s' (storage st)\) - | None \ throw Err) - } - | None \ throw Err) - } - | _ \ throw Err) - } - | (LStoreloc l, _) \ - do { - so \ applyf (\st. fmlookup (storage st) (address env)); - (case so of - Some s \ - do { - cs \ applyf (\st. cpm2s p l x t (memory st) s); - (case cs of - Some s' \ modify (\st. st \storage := fmupd (address env) s' (storage st)\) - | None \ throw Err) - } - | None \ throw Err) - } - | (LMemloc l, _) \ modify (\st. st \memory := updateStore l (MPointer p) (memory st)\) - | _ \ throw Err) - } - | (KStoptr p, Storage (STArray x t)) \ - do { - rl \ lexp lv e\<^sub>p env cd; - (case rl of - (LStackloc l, Memory _) \ - do { - sv \ applyf (\st. accessStore l (stack st)); - (case sv of - Some (KMemptr p') \ - do { - so \ applyf (\st. fmlookup (storage st) (address env)); - (case so of - Some s \ - do { - cs \ applyf (\st. cps2m p p' x t s (memory st)); - (case cs of - Some m \ modify (\st. st\memory := m\) - | None \ throw Err) - } - | None \ throw Err) - } - | _ \ throw Err) - } - | (LStackloc l, Storage _) \ modify (\st. st\stack := updateStore l (KStoptr p) (stack st)\) - | (LStoreloc l, _) \ - do { - so \ applyf (\st. fmlookup (storage st) (address env)); - (case so of - Some s \ - (case copy p l x t s of - Some s' \ modify (\st. st \storage := fmupd (address env) s' (storage st)\) - | None \ throw Err) - | None \ throw Err) - } - | (LMemloc l, _) \ - do { - so \ applyf (\st. fmlookup (storage st) (address env)); - (case so of - Some s \ - do { - cs \ applyf (\st. cps2m p l x t s (memory st)); - (case cs of - Some m \ modify (\st. st\memory := m\) - | None \ throw Err) - } - | None \ throw Err) - } - | _ \ throw Err) - } - | (KStoptr p, Storage (STMap t t')) \ - do { - rl \ lexp lv e\<^sub>p env cd; - (case rl of - (LStackloc l, _) \ modify (\st. st\stack := updateStore l (KStoptr p) (stack st)\) - | _ \ throw Err) - } - | _ \ throw Err) - }) st" -| "stmt (COMP s1 s2) e\<^sub>p e cd st = - (do { - gascheck (costs (COMP s1 s2) e\<^sub>p e cd); - stmt s1 e\<^sub>p e cd; - stmt s2 e\<^sub>p e cd - }) st" -| "stmt (ITE ex s1 s2) e\<^sub>p e cd st = - (do { - gascheck (costs (ITE ex s1 s2) e\<^sub>p e cd); - v \ expr ex e\<^sub>p e cd; - (case v of - (KValue b, Value TBool) \ - (if b = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True - then stmt s1 e\<^sub>p e cd - else stmt s2 e\<^sub>p e cd) - | _ \ throw Err) - }) st" -| "stmt (WHILE ex s0) e\<^sub>p e cd st = - (do { - gascheck (costs (WHILE ex s0) e\<^sub>p e cd); - v \ expr ex e\<^sub>p e cd; - (case v of - (KValue b, Value TBool) \ - (if b = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True - then do { - stmt s0 e\<^sub>p e cd; - stmt (WHILE ex s0) e\<^sub>p e cd - } - else return ()) - | _ \ throw Err) - }) st" -| "stmt (INVOKE i xe) e\<^sub>p e cd st = - (do { - gascheck (costs (INVOKE i xe) e\<^sub>p e cd); - (case fmlookup e\<^sub>p (address e) of - Some (ct, _) \ - (case fmlookup ct i of - Some (Method (fp, f, None)) \ - (let e' = ffold_init ct (emptyEnv (address e) (sender e) (svalue e)) (fmdom ct) - in (do { - st' \ applyf (\st. (st\stack:=emptyStore\)); - (e'', cd', st'') \ load False fp xe e\<^sub>p e' emptyStore st' e cd; - st''' \ get; - put st''; - stmt f e\<^sub>p e'' cd'; - modify (\st. st\stack:=stack st''', memory := memory st'''\) - })) - | _ \ throw Err) - | None \ throw Err) - }) st" -(*External Method calls allow to send some money val with it*) -(*However this transfer does NOT trigger a fallback*) -| "stmt (EXTERNAL ad i xe val) e\<^sub>p e cd st = - (do { - gascheck (costs (EXTERNAL ad i xe val) e\<^sub>p e cd); - kad \ expr ad e\<^sub>p e cd; - (case kad of - (KValue adv, Value TAddr) \ - (case fmlookup e\<^sub>p adv of - Some (ct, fb) \ - (do { - kv \ expr val e\<^sub>p e cd; - (case kv of - (KValue v, Value t) \ - (case fmlookup ct i of - Some (Method (fp, f, None)) \ - let e' = ffold_init ct (emptyEnv adv (address e) v) (fmdom ct) - in (do { - st' \ applyf (\st. st\stack:=emptyStore, memory:=emptyStore\); - (e'', cd', st'') \ load True fp xe e\<^sub>p e' emptyStore st' e cd; - st''' \ get; - (case transfer (address e) adv v (accounts st'') of - Some acc \ - do { - put (st''\accounts := acc\); - stmt f e\<^sub>p e'' cd'; - modify (\st. st\stack:=stack st''', memory := memory st'''\) - } - | None \ throw Err) - }) - | None \ - do { - st' \ get; - (case transfer (address e) adv v (accounts st') of - Some acc \ - do { - st'' \ get; - modify (\st. st\accounts := acc,stack:=emptyStore, memory:=emptyStore\); - stmt fb e\<^sub>p (emptyEnv adv (address e) v) cd; - modify (\st. st\stack:=stack st'', memory := memory st''\) - } - | None \ throw Err) - } - | _ \ throw Err) - | _ \ throw Err) - }) - | None \ throw Err) - | _ \ throw Err) - }) st" -| "stmt (TRANSFER ad ex) e\<^sub>p e cd st = - (do { - gascheck (costs (TRANSFER ad ex) e\<^sub>p e cd); - kv \ expr ex e\<^sub>p e cd; - (case kv of - (KValue v, Value t) \ - (do { - kv' \ expr ad e\<^sub>p e cd; - (case kv' of - (KValue adv, Value TAddr) \ - (do { - acs \ applyf accounts; - (case transfer (address e) adv v acs of - Some acc \ (case fmlookup e\<^sub>p adv of - Some (ct, f) \ - let e' = ffold_init ct (emptyEnv adv (address e) v) (fmdom ct) - in (do { - st' \ get; - modify (\st. (st\accounts := acc, stack:=emptyStore, memory:=emptyStore\)); - stmt f e\<^sub>p e' emptyStore; - modify (\st. st\stack:=stack st', memory := memory st'\) - }) - | None \ modify (\st. (st\accounts := acc\))) - | None \ throw Err) - }) - | _ \ throw Err) - }) - | _ \ throw Err) - }) st" -| "stmt (BLOCK ((id0, tp), ex) s) e\<^sub>p e\<^sub>v cd st = - (do { - gascheck (costs (BLOCK ((id0, tp), ex) s) e\<^sub>p e\<^sub>v cd); - (case ex of - None \ (do { - mem \ applyf memory; - (cd', e') \ decl id0 tp None False cd mem cd e\<^sub>v; - stmt s e\<^sub>p e' cd' - }) - | Some ex' \ (do { - (v, t) \ expr ex' e\<^sub>p e\<^sub>v cd; - mem \ applyf memory; - (cd', e') \ decl id0 tp (Some (v, t)) False cd mem cd e\<^sub>v; - stmt s e\<^sub>p e' cd' - })) - }) st" - by pat_completeness auto - -subsection \Gas Consumption\ - -lemma lift_gas: - assumes "lift expr f e1 e2 e\<^sub>p e cd st = Normal ((v, t), st4')" - and "\st4' v4 t4. expr e1 e\<^sub>p e cd st = Normal ((v4, t4), st4') \ gas st4' \ gas st" - and "\x1 x y xa ya x1a x1b st4' v4 t4. expr e1 e\<^sub>p e cd st = Normal (x, y) - \ (xa, ya) = x - \ xa = KValue x1a - \ ya = Value x1b - \ expr e2 e\<^sub>p e cd y = Normal ((v4, t4), st4') - \ gas st4' \ gas y" - shows "gas st4' \ gas st" -proof (cases "expr e1 e\<^sub>p e cd st") - case (n a st') - then show ?thesis - proof (cases a) - case (Pair b c) - then show ?thesis - proof (cases b) - case (KValue v1) - then show ?thesis - proof (cases c) - case (Value t1) - then show ?thesis - proof (cases "expr e2 e\<^sub>p e cd st'") - case r2: (n a' st'') - then show ?thesis - proof (cases a') - case p2: (Pair b c) - then show ?thesis - proof (cases b) - case v2: (KValue v2) - then show ?thesis - proof (cases c) - case t2: (Value t2) - then show ?thesis - proof (cases "f t1 t2 v1 v2") - case None - with assms n Pair KValue Value r2 p2 v2 t2 show ?thesis by simp - next - case (Some a'') - then show ?thesis - proof (cases a'') - case p3: (Pair v t) - with assms n Pair KValue Value r2 p2 v2 t2 Some have "gas st4'\gas st''" by simp - moreover from assms n Pair KValue Value r2 p2 v2 t2 Some have "gas st''\gas st'" by simp - moreover from assms n Pair KValue Value r2 p2 v2 t2 Some have "gas st'\gas st" by simp - ultimately show ?thesis by arith - qed - qed - next - case (Calldata x2) - with assms n Pair KValue Value r2 p2 v2 show ?thesis by simp - next - case (Memory x3) - with assms n Pair KValue Value r2 p2 v2 show ?thesis by simp - next - case (Storage x4) - with assms n Pair KValue Value r2 p2 v2 show ?thesis by simp - qed - next - case (KCDptr x2) - with assms n Pair KValue Value r2 p2 show ?thesis by simp - next - case (KMemptr x3) - with assms n Pair KValue Value r2 p2 show ?thesis by simp - next - case (KStoptr x4) - with assms n Pair KValue Value r2 p2 show ?thesis by simp - qed - qed - next - case (e x) - with assms n Pair KValue Value show ?thesis by simp - qed - next - case (Calldata x2) - with assms n Pair KValue show ?thesis by simp - next - case (Memory x3) - with assms n Pair KValue show ?thesis by simp - next - case (Storage x4) - with assms n Pair KValue show ?thesis by simp - qed - next - case (KCDptr x2) - with assms n Pair show ?thesis by simp - next - case (KMemptr x3) - with assms n Pair show ?thesis by simp - next - case (KStoptr x4) - with assms n Pair show ?thesis by simp - qed - qed -next - case (e x) - with assms show ?thesis by simp -qed - -lemma msel_ssel_lexp_expr_load_rexp_stmt_dom_gas: - "msel_ssel_lexp_expr_load_rexp_stmt_dom (Inl (Inl (c1, t1, l1, xe1, ep1, ev1, cd1, st1))) - \ (\l1' t1' st1'. msel c1 t1 l1 xe1 ep1 ev1 cd1 st1 = Normal ((l1', t1'), st1') \ gas st1' \ gas st1)" - "msel_ssel_lexp_expr_load_rexp_stmt_dom (Inl (Inr (Inl (t2, l2, xe2, ep2, ev2, cd2, st2)))) - \ (\l2' t2' st2'. ssel t2 l2 xe2 ep2 ev2 cd2 st2 = Normal ((l2', t2'), st2') \ gas st2' \ gas st2)" - "msel_ssel_lexp_expr_load_rexp_stmt_dom (Inl (Inr (Inr (l5, ep5, ev5, cd5, st5)))) - \ (\l5' t5' st5'. lexp l5 ep5 ev5 cd5 st5 = Normal ((l5', t5'), st5') \ gas st5' \ gas st5)" - "msel_ssel_lexp_expr_load_rexp_stmt_dom (Inr (Inl (Inl (e4, ep4, ev4, cd4, st4)))) - \ (\st4' v4 t4. expr e4 ep4 ev4 cd4 st4 = Normal ((v4, t4), st4') \ gas st4' \ gas st4)" - "msel_ssel_lexp_expr_load_rexp_stmt_dom (Inr (Inl (Inr (lcp, lis, lxs, lep, lev0, lcd0, lst0, lev, lcd, lst)))) - \ (\ev cd st st'. load lcp lis lxs lep lev0 lcd0 lst0 lev lcd lst = Normal ((ev, cd, st), st') \ gas st \ gas lst0 \ gas st' \ gas lst \ address ev = address lev0)" - "msel_ssel_lexp_expr_load_rexp_stmt_dom (Inr (Inr (Inl (l3, ep3, ev3, cd3, st3)))) - \ (\l3' t3' st3'. rexp l3 ep3 ev3 cd3 st3 = Normal ((l3', t3'), st3') \ gas st3' \ gas st3)" - "msel_ssel_lexp_expr_load_rexp_stmt_dom (Inr (Inr (Inr (s6, ep6, ev6, cd6, st6)))) - \ (\st6'. stmt s6 ep6 ev6 cd6 st6 = Normal((), st6') \ gas st6' \ gas st6)" -proof (induct rule: msel_ssel_lexp_expr_load_rexp_stmt.pinduct -[where ?P1.0="\c1 t1 l1 xe1 ep1 ev1 cd1 st1. (\l1' t1' st1'. msel c1 t1 l1 xe1 ep1 ev1 cd1 st1 = Normal ((l1', t1'), st1') \ gas st1' \ gas st1)" - and ?P2.0="\t2 l2 xe2 ep2 ev2 cd2 st2. (\l2' t2' st2'. ssel t2 l2 xe2 ep2 ev2 cd2 st2 = Normal ((l2', t2'), st2') \ gas st2' \ gas st2)" - and ?P3.0="\l5 ep5 ev5 cd5 st5. (\l5' t5' st5'. lexp l5 ep5 ev5 cd5 st5 = Normal ((l5', t5'), st5') \ gas st5' \ gas st5)" - and ?P4.0="\e4 ep4 ev4 cd4 st4. (\st4' v4 t4. expr e4 ep4 ev4 cd4 st4 = Normal ((v4, t4), st4') \ gas st4' \ gas st4)" - and ?P5.0="\lcp lis lxs lep lev0 lcd0 lst0 lev lcd lst. (\ev cd st st'. load lcp lis lxs lep lev0 lcd0 lst0 lev lcd lst = Normal ((ev, cd, st), st') \ gas st \ gas lst0 \ gas st' \ gas lst \ address ev = address lev0)" - and ?P6.0="\l3 ep3 ev3 cd3 st3. (\l3' t3' st3'. rexp l3 ep3 ev3 cd3 st3 = Normal ((l3', t3'), st3') \ gas st3' \ gas st3)" - and ?P7.0="\s6 ep6 ev6 cd6 st6. (\st6'. stmt s6 ep6 ev6 cd6 st6 = Normal ((), st6') \ gas st6' \ gas st6)" -]) - case (1 uu uv uw ux uy uz va) - then show ?case using msel.psimps(1) by auto -next - case (2 vb vc vd ve vf vg vh vi) - then show ?case using msel.psimps(2) by auto -next - case (3 vj al t loc x e\<^sub>p env cd st) - then show ?case using msel.psimps(3) by (auto split: if_split_asm Type.split_asm Stackvalue.split_asm prod.split_asm StateMonad.result.split_asm) -next - case (4 mm al t loc x y ys e\<^sub>p env cd st) - show ?case - proof (rule allI[THEN allI, THEN allI, OF impI]) - fix l1' t1' st1' assume a1: "msel mm (MTArray al t) loc (x # y # ys) e\<^sub>p env cd st = Normal ((l1', t1'), st1')" - show "gas st1' \ gas st" - proof (cases "expr x e\<^sub>p env cd st") - case (n a st') - then show ?thesis - proof (cases a) - case (Pair b c) - then show ?thesis - proof (cases b) - case (KValue v) - then show ?thesis - proof (cases c) - case (Value t') - then show ?thesis - proof (cases) - assume l: "less t' (TUInt 256) v (ShowL\<^sub>i\<^sub>n\<^sub>t al) = Some (ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True, TBool)" - then show ?thesis - proof (cases "accessStore (hash loc v) (if mm then memory st' else cd)") - case None - with 4 a1 n Pair KValue Value l show ?thesis using msel.psimps(4) by simp - next - case (Some a) - then show ?thesis - proof (cases a) - case (MValue x1) - with 4 a1 n Pair KValue Value Some l show ?thesis using msel.psimps(4) by simp - next - case (MPointer l) - with n Pair KValue Value l Some - have "msel mm (MTArray al t) loc (x # y # ys) e\<^sub>p env cd st = msel mm t l (y # ys) e\<^sub>p env cd st'" - using msel.psimps(4) 4(1) by simp - moreover from n Pair have "gas st' \ gas st" using 4(2) by simp - moreover from a1 MPointer n Pair KValue Value l Some - have "gas st1' \ gas st'" using msel.psimps(4) 4(3) 4(1) by simp - ultimately show ?thesis by simp - qed - qed - next - assume "\ less t' (TUInt 256) v (ShowL\<^sub>i\<^sub>n\<^sub>t al) = Some (ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True, TBool)" - with 4 a1 n Pair KValue Value show ?thesis using msel.psimps(4) by simp - qed - next - case (Calldata x2) - with 4 a1 n Pair KValue show ?thesis using msel.psimps(4) by simp - next - case (Memory x3) - with 4 a1 n Pair KValue show ?thesis using msel.psimps(4) by simp - next - case (Storage x4) - with 4 a1 n Pair KValue show ?thesis using msel.psimps(4) by simp - qed - next - case (KCDptr x2) - with 4 a1 n Pair show ?thesis using msel.psimps(4) by simp - next - case (KMemptr x3) - with 4 a1 n Pair show ?thesis using msel.psimps(4) by simp - next - case (KStoptr x4) - with 4 a1 n Pair show ?thesis using msel.psimps(4) by simp - qed - qed - next - case (e x) - with 4 a1 show ?thesis using msel.psimps(4) by simp - qed - qed -next - case (5 tp loc vk vl vm st) - then show ?case using ssel.psimps(1) by auto -next - case (6 vn vo vp vq vr vs vt vu) - then show ?case using ssel.psimps(2) by auto -next - case (7 al t loc x xs e\<^sub>p env cd st) - show ?case - proof (rule allI[THEN allI, THEN allI, OF impI]) - fix l2' t2' st2' assume a1: "ssel (STArray al t) loc (x # xs) e\<^sub>p env cd st = Normal ((l2', t2'), st2')" - show "gas st2' \ gas st" - proof (cases "expr x e\<^sub>p env cd st") - case (n a st'') - then show ?thesis - proof (cases a) - case (Pair b c) - then show ?thesis - proof (cases b) - case (KValue v) - then show ?thesis - proof (cases c) - case (Value t') - then show ?thesis - proof (cases) - assume l: "less t' (TUInt 256) v (ShowL\<^sub>i\<^sub>n\<^sub>t al) = Some (ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True, TBool)" - with n Pair KValue Value l - have "ssel (STArray al t) loc (x # xs) e\<^sub>p env cd st = ssel t (hash loc v) xs e\<^sub>p env cd st''" - using ssel.psimps(3) 7(1) by simp - moreover from n Pair have "gas st'' \ gas st" using 7(2) by simp - moreover from a1 n Pair KValue Value l - have "gas st2' \ gas st''" using ssel.psimps(3) 7(3) 7(1) by simp - ultimately show ?thesis by simp - next - assume "\ less t' (TUInt 256) v (ShowL\<^sub>i\<^sub>n\<^sub>t al) = Some (ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True, TBool)" - with 7 a1 n Pair KValue Value show ?thesis using ssel.psimps(3) by simp - qed - next - case (Calldata x2) - with 7 a1 n Pair KValue show ?thesis using ssel.psimps(3) by simp - next - case (Memory x3) - with 7 a1 n Pair KValue show ?thesis using ssel.psimps(3) by simp - next - case (Storage x4) - with 7 a1 n Pair KValue show ?thesis using ssel.psimps(3) by simp - qed - next - case (KCDptr x2) - with 7 a1 n Pair show ?thesis using ssel.psimps(3) by simp - next - case (KMemptr x3) - with 7 a1 n Pair show ?thesis using ssel.psimps(3) by simp - next - case (KStoptr x4) - with 7 a1 n Pair show ?thesis using ssel.psimps(3) by simp - qed - qed - next - case (e e) - with 7 a1 show ?thesis using ssel.psimps(3) by simp - qed - qed -next - case (8 vv t loc x xs e\<^sub>p env cd st) - show ?case - proof (rule allI[THEN allI, THEN allI, OF impI]) - fix l2' t2' st2' assume a1: "ssel (STMap vv t) loc (x # xs) e\<^sub>p env cd st = Normal ((l2', t2'), st2')" - show "gas st2' \ gas st" - proof (cases "expr x e\<^sub>p env cd st") - case (n a st') - then show ?thesis - proof (cases a) - case (Pair b c) - then show ?thesis - proof (cases b) - case (KValue v) - with 8 n Pair have "ssel (STMap vv t) loc (x # xs) e\<^sub>p env cd st = ssel t (hash loc v) xs e\<^sub>p env cd st'" using ssel.psimps(4) by simp - moreover from n Pair have "gas st' \ gas st" using 8(2) by simp - moreover from a1 n Pair KValue - have "gas st2' \ gas st'" using ssel.psimps(4) 8(3) 8(1) by simp - ultimately show ?thesis by simp - next - case (KCDptr x2) - with 8 a1 n Pair show ?thesis using ssel.psimps(4) by simp - next - case (KMemptr x3) - with 8 a1 n Pair show ?thesis using ssel.psimps(4) by simp - next - case (KStoptr x4) - with 8 a1 n Pair show ?thesis using ssel.psimps(4) by simp - qed - qed - next - case (e x) - with 8 a1 show ?thesis using ssel.psimps(4) by simp - qed - qed -next - case (9 i vw e vx st) - then show ?case using lexp.psimps(1)[of i vw e vx st] by (simp split: option.split_asm Denvalue.split_asm prod.split_asm) -next - case (10 i r e\<^sub>p e cd st) - show ?case - proof (rule allI[THEN allI, THEN allI, OF impI]) - fix st5' xa xaa - assume a1: "lexp (Ref i r) e\<^sub>p e cd st = Normal ((st5', xa), xaa)" - then show "gas xaa \ gas st" - proof (cases "fmlookup (denvalue e) i") - case None - with 10 a1 show ?thesis using lexp.psimps(2) by simp - next - case (Some a) - then show ?thesis - proof (cases a) - case (Pair tp b) - then show ?thesis - proof (cases b) - case (Stackloc l) - then show ?thesis - proof (cases "accessStore l (stack st)") - case None - with 10 a1 Some Pair Stackloc show ?thesis using lexp.psimps(2) by simp - next - case s2: (Some a) - then show ?thesis - proof (cases a) - case (KValue x1) - with 10 a1 Some Pair Stackloc s2 show ?thesis using lexp.psimps(2) by simp - next - case (KCDptr x2) - with 10 a1 Some Pair Stackloc s2 show ?thesis using lexp.psimps(2) by simp - next - case (KMemptr l') - then show ?thesis - proof (cases tp) - case (Value x1) - with 10 a1 Some Pair Stackloc s2 KMemptr show ?thesis using lexp.psimps(2) by simp - next - case (Calldata x2) - with 10 a1 Some Pair Stackloc s2 KMemptr show ?thesis using lexp.psimps(2) by simp - next - case (Memory t) - then show ?thesis - proof (cases "msel True t l' r e\<^sub>p e cd st") - case (n a s) - with 10 a1 Some Pair Stackloc s2 KMemptr Memory show ?thesis using lexp.psimps(2) by (simp split: prod.split_asm) - next - case (e e) - with 10 a1 Some Pair Stackloc s2 KMemptr Memory show ?thesis using lexp.psimps(2) by simp - qed - next - case (Storage x4) - with 10 a1 Some Pair Stackloc s2 KMemptr show ?thesis using lexp.psimps(2) by simp - qed - next - case (KStoptr l') - then show ?thesis - proof (cases tp) - case (Value x1) - with 10 a1 Some Pair Stackloc s2 KStoptr show ?thesis using lexp.psimps(2) by simp - next - case (Calldata x2) - with 10 a1 Some Pair Stackloc s2 KStoptr show ?thesis using lexp.psimps(2) by simp - next - case (Memory t) - with 10 a1 Some Pair Stackloc s2 KStoptr show ?thesis using lexp.psimps(2) by simp - next - case (Storage t) - then show ?thesis - proof (cases "ssel t l' r e\<^sub>p e cd st") - case (n a s) - with 10 a1 Some Pair Stackloc s2 KStoptr Storage show ?thesis using lexp.psimps(2) by (auto split: prod.split_asm) - next - case (e x) - with 10 a1 Some Pair Stackloc s2 KStoptr Storage show ?thesis using lexp.psimps(2) by simp - qed - qed - qed - qed - next - case (Storeloc l) - then show ?thesis - proof (cases tp) - case (Value x1) - with 10 a1 Some Pair Storeloc show ?thesis using lexp.psimps(2) by simp - next - case (Calldata x2) - with 10 a1 Some Pair Storeloc show ?thesis using lexp.psimps(2) by simp - next - case (Memory t) - with 10 a1 Some Pair Storeloc show ?thesis using lexp.psimps(2) by simp - next - case (Storage t) - then show ?thesis - proof (cases "ssel t l r e\<^sub>p e cd st") - case (n a s) - with 10 a1 Some Pair Storeloc Storage show ?thesis using lexp.psimps(2) by (auto split: prod.split_asm) - next - case (e x) - with 10 a1 Some Pair Storeloc Storage show ?thesis using lexp.psimps(2) by simp - qed - qed - qed - qed - qed - qed -next - case (11 b x e\<^sub>p e vy st) - then show ?case using expr.psimps(1) by (simp split:if_split_asm) -next - case (12 b x e\<^sub>p e vz st) - then show ?case using expr.psimps(2) by (simp split:if_split_asm) -next - case (13 ad e\<^sub>p e wa st) - then show ?case using expr.psimps(3) by simp -next - case (14 ad e\<^sub>p e wb st) - define g where "g = costs\<^sub>e (BALANCE ad) e\<^sub>p e wb st" - show ?case - proof (rule allI[THEN allI, THEN allI, OF impI]) - fix t4 xa xaa - assume *: "expr (BALANCE ad) e\<^sub>p e wb st = Normal ((xa, xaa), t4)" - show "gas t4 \ gas st" - proof (cases) - assume "gas st \ g" - with 14 g_def * show ?thesis using expr.psimps(4) by simp - next - assume gcost: "\ gas st \ g" - then show ?thesis - proof (cases "expr ad e\<^sub>p e wb (st\gas := gas st - g\)") - case (n a s) - show ?thesis - proof (cases a) - case (Pair b c) - then show ?thesis - proof (cases b) - case (KValue x1) - then show ?thesis - proof (cases c) - case (Value x1) - then show ?thesis - proof (cases x1) - case (TSInt x1) - with 14 g_def * gcost n Pair KValue Value show ?thesis using expr.psimps(4)[of ad e\<^sub>p e wb st] by simp - next - case (TUInt x2) - with 14 g_def * gcost n Pair KValue Value show ?thesis using expr.psimps(4)[of ad e\<^sub>p e wb st] by simp - next - case TBool - with 14 g_def * gcost n Pair KValue Value show ?thesis using expr.psimps(4)[of ad e\<^sub>p e wb st] by simp - next - case TAddr - with 14 g_def * gcost n Pair KValue Value show "gas t4 \ gas st" using expr.psimps(4)[of ad e\<^sub>p e wb st] by simp - qed - next - case (Calldata x2) - with 14 g_def * gcost n Pair KValue show ?thesis using expr.psimps(4)[of ad e\<^sub>p e wb st] by simp - next - case (Memory x3) - with 14 g_def * gcost n Pair KValue show ?thesis using expr.psimps(4)[of ad e\<^sub>p e wb st] by simp - next - case (Storage x4) - with 14 g_def * gcost n Pair KValue show ?thesis using expr.psimps(4)[of ad e\<^sub>p e wb st] by simp - qed - next - case (KCDptr x2) - with 14 g_def * gcost n Pair show ?thesis using expr.psimps(4)[of ad e\<^sub>p e wb st] by simp - next - case (KMemptr x3) - with 14 g_def * gcost n Pair show ?thesis using expr.psimps(4)[of ad e\<^sub>p e wb st] by simp - next - case (KStoptr x4) - with 14 g_def * gcost n Pair show ?thesis using expr.psimps(4)[of ad e\<^sub>p e wb st] by simp - qed - qed - next - case (e _) - with 14 g_def * gcost show ?thesis using expr.psimps(4)[of ad e\<^sub>p e wb st] by simp - qed - qed - qed -next - case (15 e\<^sub>p e wc st) - then show ?case using expr.psimps(5) by simp -next - case (16 e\<^sub>p e wd st) - then show ?case using expr.psimps(6) by simp -next - case (17 e\<^sub>p e wd st) - then show ?case using expr.psimps(7) by simp -next - case (18 e\<^sub>p e wd st) - then show ?case using expr.psimps(8) by simp -next - case (19 e\<^sub>p e wd st) - then show ?case using expr.psimps(9) by simp -next - case (20 x e\<^sub>p e cd st) - define g where "g = costs\<^sub>e (NOT x) e\<^sub>p e cd st" - show ?case - proof (rule allI[THEN allI, THEN allI, OF impI]) - fix st4' v4 t4 assume a1: "expr (NOT x) e\<^sub>p e cd st = Normal ((v4, t4), st4')" - show "gas st4' \ gas st" - proof (cases) - assume "gas st \ g" - with 20 g_def a1 show ?thesis using expr.psimps by simp - next - assume gcost: "\ gas st \ g" - then show ?thesis - proof (cases "expr x e\<^sub>p e cd (st\gas := gas st - g\)") - case (n a st') - then show ?thesis - proof (cases a) - case (Pair b c) - then show ?thesis - proof (cases b) - case (KValue v) - then show ?thesis - proof (cases c) - case (Value t) - then show ?thesis - proof (cases) - assume v_def: "v = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True" - with 20(1) g_def gcost n Pair KValue Value have "expr (NOT x) e\<^sub>p e cd st = expr FALSE e\<^sub>p e cd st'" using expr.psimps(10) by simp - moreover from 20(2) g_def gcost n Pair have "gas st' \ gas st" by simp - moreover from 20(1) 20(3) a1 g_def gcost n Pair KValue Value v_def have "gas st4' \ gas st'" using expr.psimps(10) by simp - ultimately show ?thesis by arith - next - assume v_def: "\ v = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True" - then show ?thesis - proof (cases) - assume v_def2: "v = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l False" - with 20(1) g_def gcost n Pair KValue Value v_def have "expr (NOT x) e\<^sub>p e cd st = expr TRUE e\<^sub>p e cd st'" using expr.psimps(10) by simp - moreover from 20(2) g_def gcost n Pair have "gas st' \ gas st" by simp - moreover from 20 a1 g_def gcost n Pair KValue Value v_def v_def2 have "gas st4' \ gas st'" using expr.psimps(10) by simp - ultimately show ?thesis by arith - next - assume "\ v = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l False" - with 20 a1 g_def gcost n Pair KValue Value v_def show ?thesis using expr.psimps(10) by simp - qed - qed - next - case (Calldata x2) - with 20 a1 g_def gcost n Pair KValue show ?thesis using expr.psimps(10) by simp - next - case (Memory x3) - with 20 a1 g_def gcost n Pair KValue show ?thesis using expr.psimps(10) by simp - next - case (Storage x4) - with 20 a1 g_def gcost n Pair KValue show ?thesis using expr.psimps(10) by simp - qed - next - case (KCDptr x2) - with 20 a1 g_def gcost n Pair show ?thesis using expr.psimps(10) by simp - next - case (KMemptr x3) - with 20 a1 g_def gcost n Pair show ?thesis using expr.psimps(10) by simp - next - case (KStoptr x4) - with 20 a1 g_def gcost n Pair show ?thesis using expr.psimps(10) by simp - qed - qed - next - case (e e) - with 20 a1 g_def gcost show ?thesis using expr.psimps(10) by simp - qed - qed - qed -next - case (21 e1 e2 e\<^sub>p e cd st) - define g where "g = costs\<^sub>e (PLUS e1 e2) e\<^sub>p e cd st" - show ?case - proof (rule allI[THEN allI, THEN allI, OF impI]) - fix t4 xa xaa assume e_def: "expr (PLUS e1 e2) e\<^sub>p e cd st = Normal ((xa, xaa), t4)" - then show "gas t4 \ gas st" - proof (cases) - assume "gas st \ g" - with 21(1) e_def show ?thesis using expr.psimps(11) g_def by simp - next - assume "\ gas st \ g" - with 21(1) e_def g_def have "lift expr add e1 e2 e\<^sub>p e cd (st\gas := gas st - g\) = Normal ((xa, xaa), t4)" using expr.psimps(11)[of e1 e2 e\<^sub>p e cd st] by simp - moreover from 21(2) `\ gas st \ g` g_def have "(\st4' v4 t4. expr e1 e\<^sub>p e cd (st\gas := gas st - g\) = Normal ((v4, t4), st4') \ gas st4' \ gas (st\gas := gas st - g\))" by simp - moreover from 21(3) `\ gas st \ g` g_def have "(\x1 x y xa ya x1a x1b st4' v4 t4. - expr e1 e\<^sub>p e cd (st\gas := gas st - g\) = Normal (x, y) \ - (xa, ya) = x \ - xa = KValue x1a \ - ya = Value x1b \ expr e2 e\<^sub>p e cd y = Normal ((v4, t4), st4') \ gas st4' \ gas y)" by auto - ultimately show "gas t4 \ gas st" using lift_gas[of e1 e\<^sub>p e cd e2 "add" "st\gas := gas st - g\" xa xaa t4] by simp - qed - qed -next - case (22 e1 e2 e\<^sub>p e cd st) - define g where "g = costs\<^sub>e (MINUS e1 e2) e\<^sub>p e cd st" - show ?case - proof (rule allI[THEN allI, THEN allI, OF impI]) - fix t4 xa xaa assume e_def: "expr (MINUS e1 e2) e\<^sub>p e cd st = Normal ((xa, xaa), t4)" - then show "gas t4 \ gas st" - proof (cases) - assume "gas st \ g" - with 22(1) e_def show ?thesis using expr.psimps(12) g_def by simp - next - assume "\ gas st \ g" - with 22(1) e_def g_def have "lift expr sub e1 e2 e\<^sub>p e cd (st\gas := gas st - g\) = Normal ((xa, xaa), t4)" using expr.psimps(12)[of e1 e2 e\<^sub>p e cd st] by simp - moreover from 22(2) `\ gas st \ g` g_def have "(\st4' v4 t4. expr e1 e\<^sub>p e cd (st\gas := gas st - g\) = Normal ((v4, t4), st4') \ gas st4' \ gas (st\gas := gas st - g\))" by simp - moreover from 22(3) `\ gas st \ g` g_def have "(\x1 x y xa ya x1a x1b st4' v4 t4. - expr e1 e\<^sub>p e cd (st\gas := gas st - g\) = Normal (x, y) \ - (xa, ya) = x \ - xa = KValue x1a \ - ya = Value x1b \ expr e2 e\<^sub>p e cd y = Normal ((v4, t4), st4') \ gas st4' \ gas y)" by auto - ultimately show "gas t4 \ gas st" using lift_gas[of e1 e\<^sub>p e cd e2 "sub" "st\gas := gas st - g\" xa xaa t4] by simp - qed - qed -next - case (23 e1 e2 e\<^sub>p e cd st) - define g where "g = costs\<^sub>e (LESS e1 e2) e\<^sub>p e cd st" - show ?case - proof (rule allI[THEN allI, THEN allI, OF impI]) - fix t4 xa xaa assume e_def: "expr (LESS e1 e2) e\<^sub>p e cd st = Normal ((xa, xaa), t4)" - then show "gas t4 \ gas st" - proof (cases) - assume "gas st \ g" - with 23(1) e_def show ?thesis using expr.psimps(13) g_def by simp - next - assume "\ gas st \ g" - with 23(1) e_def g_def have "lift expr less e1 e2 e\<^sub>p e cd (st\gas := gas st - g\) = Normal ((xa, xaa), t4)" using expr.psimps(13)[of e1 e2 e\<^sub>p e cd st] by simp - moreover from 23(2) `\ gas st \ g` g_def have "(\st4' v4 t4. expr e1 e\<^sub>p e cd (st\gas := gas st - g\) = Normal ((v4, t4), st4') \ gas st4' \ gas (st\gas := gas st - g\))" by simp - moreover from 23(3) `\ gas st \ g` g_def have "(\x1 x y xa ya x1a x1b st4' v4 t4. - expr e1 e\<^sub>p e cd (st\gas := gas st - g\) = Normal (x, y) \ - (xa, ya) = x \ - xa = KValue x1a \ - ya = Value x1b \ expr e2 e\<^sub>p e cd y = Normal ((v4, t4), st4') \ gas st4' \ gas y)" by auto - ultimately show "gas t4 \ gas st" using lift_gas[of e1 e\<^sub>p e cd e2 "less" "st\gas := gas st - g\" xa xaa t4] by simp - qed - qed -next - case (24 e1 e2 e\<^sub>p e cd st) - define g where "g = costs\<^sub>e (EQUAL e1 e2) e\<^sub>p e cd st" - show ?case - proof (rule allI[THEN allI, THEN allI, OF impI]) - fix t4 xa xaa assume e_def: "expr (EQUAL e1 e2) e\<^sub>p e cd st = Normal ((xa, xaa), t4)" - then show "gas t4 \ gas st" - proof (cases) - assume "gas st \ g" - with 24(1) e_def show ?thesis using expr.psimps(14) g_def by simp - next - assume "\ gas st \ g" - with 24(1) e_def g_def have "lift expr equal e1 e2 e\<^sub>p e cd (st\gas := gas st - g\) = Normal ((xa, xaa), t4)" using expr.psimps(14)[of e1 e2 e\<^sub>p e cd st] by simp - moreover from 24(2) `\ gas st \ g` g_def have "(\st4' v4 t4. expr e1 e\<^sub>p e cd (st\gas := gas st - g\) = Normal ((v4, t4), st4') \ gas st4' \ gas (st\gas := gas st - g\))" by simp - moreover from 24(3) `\ gas st \ g` g_def have "(\x1 x y xa ya x1a x1b st4' v4 t4. - expr e1 e\<^sub>p e cd (st\gas := gas st - g\) = Normal (x, y) \ - (xa, ya) = x \ - xa = KValue x1a \ - ya = Value x1b \ expr e2 e\<^sub>p e cd y = Normal ((v4, t4), st4') \ gas st4' \ gas y)" by auto - ultimately show "gas t4 \ gas st" using lift_gas[of e1 e\<^sub>p e cd e2 "equal" "st\gas := gas st - g\" xa xaa t4] by simp - qed - qed -next - case (25 e1 e2 e\<^sub>p e cd st) - define g where "g = costs\<^sub>e (AND e1 e2) e\<^sub>p e cd st" - show ?case - proof (rule allI[THEN allI, THEN allI, OF impI]) - fix t4 xa xaa assume e_def: "expr (AND e1 e2) e\<^sub>p e cd st = Normal ((xa, xaa), t4)" - then show "gas t4 \ gas st" - proof (cases) - assume "gas st \ g" - with 25(1) e_def show ?thesis using expr.psimps(15) g_def by simp - next - assume "\ gas st \ g" - with 25(1) e_def g_def have "lift expr vtand e1 e2 e\<^sub>p e cd (st\gas := gas st - g\) = Normal ((xa, xaa), t4)" using expr.psimps(15)[of e1 e2 e\<^sub>p e cd st] by simp - moreover from 25(2) `\ gas st \ g` g_def have "(\st4' v4 t4. expr e1 e\<^sub>p e cd (st\gas := gas st - g\) = Normal ((v4, t4), st4') \ gas st4' \ gas (st\gas := gas st - g\))" by simp - moreover from 25(3) `\ gas st \ g` g_def have "(\x1 x y xa ya x1a x1b st4' v4 t4. - expr e1 e\<^sub>p e cd (st\gas := gas st - g\) = Normal (x, y) \ - (xa, ya) = x \ - xa = KValue x1a \ - ya = Value x1b \ expr e2 e\<^sub>p e cd y = Normal ((v4, t4), st4') \ gas st4' \ gas y)" by auto - ultimately show "gas t4 \ gas st" using lift_gas[of e1 e\<^sub>p e cd e2 "vtand" "st\gas := gas st - g\" xa xaa t4] by simp - qed - qed -next - case (26 e1 e2 e\<^sub>p e cd st) - define g where "g = costs\<^sub>e (OR e1 e2) e\<^sub>p e cd st" - show ?case - proof (rule allI[THEN allI, THEN allI, OF impI]) - fix t4 xa xaa assume e_def: "expr (OR e1 e2) e\<^sub>p e cd st = Normal ((xa, xaa), t4)" - then show "gas t4 \ gas st" - proof (cases) - assume "gas st \ g" - with 26(1) e_def show ?thesis using expr.psimps(16) g_def by simp - next - assume "\ gas st \ g" - with 26(1) e_def g_def have "lift expr vtor e1 e2 e\<^sub>p e cd (st\gas := gas st - g\) = Normal ((xa, xaa), t4)" using expr.psimps(16)[of e1 e2 e\<^sub>p e cd st] by simp - moreover from 26(2) `\ gas st \ g` g_def have "(\st4' v4 t4. expr e1 e\<^sub>p e cd (st\gas := gas st - g\) = Normal ((v4, t4), st4') \ gas st4' \ gas (st\gas := gas st - g\))" by simp - moreover from 26(3) `\ gas st \ g` g_def have "(\x1 x y xa ya x1a x1b st4' v4 t4. - expr e1 e\<^sub>p e cd (st\gas := gas st - g\) = Normal (x, y) \ - (xa, ya) = x \ - xa = KValue x1a \ - ya = Value x1b \ expr e2 e\<^sub>p e cd y = Normal ((v4, t4), st4') \ gas st4' \ gas y)" by auto - ultimately show "gas t4 \ gas st" using lift_gas[of e1 e\<^sub>p e cd e2 "vtor" "st\gas := gas st - g\" xa xaa t4] by simp - qed - qed -next - case (27 i e\<^sub>p e cd st) - then show ?case using expr.psimps(17) by (auto split: prod.split_asm option.split_asm StateMonad.result.split_asm) -next - case (28 i xe e\<^sub>p e cd st) - define g where "g = costs\<^sub>e (CALL i xe) e\<^sub>p e cd st" - show ?case - proof (rule allI[THEN allI, THEN allI, OF impI]) - fix st4' v4 t4 assume a1: "expr (CALL i xe) e\<^sub>p e cd st = Normal ((v4, t4), st4')" - show "gas st4' \ gas st" - proof (cases) - assume "gas st \ g" - with 28 g_def a1 show ?thesis using expr.psimps by simp - next - assume gcost: "\ gas st \ g" - then show ?thesis - proof (cases "fmlookup e\<^sub>p (address e)") - case None - with 28(1) a1 g_def gcost show ?thesis using expr.psimps(18) by simp - next - case (Some a) - then show ?thesis - proof (cases a) - case (Pair ct _) - then show ?thesis - proof (cases "fmlookup ct i") - case None - with 28(1) a1 g_def gcost Some Pair show ?thesis using expr.psimps(18) by simp - next - case s1: (Some a) - then show ?thesis - proof (cases a) - case (Method x1) - then show ?thesis - proof (cases x1) - case (fields fp f c) - then show ?thesis - proof (cases c) - case None - with 28(1) a1 g_def gcost Some Pair s1 Method fields show ?thesis using expr.psimps(18) by simp - next - case s2: (Some x) - define st' e' - where "st' = st\gas := gas st - g\\stack:=emptyStore\" - and "e' = ffold (init ct) (emptyEnv (address e) (sender e) (svalue e)) (fmdom ct)" - then show ?thesis - proof (cases "load False fp xe e\<^sub>p e' emptyStore st' e cd (st\gas := gas st - g\)") - case s4: (n a st''') - then show ?thesis - proof (cases a) - case f2: (fields e'' cd' st'') - then show ?thesis - proof (cases "stmt f e\<^sub>p e'' cd' st''") - case n2: (n a st'''') - then show ?thesis - proof (cases "expr x e\<^sub>p e'' cd' st''''") - case n3: (n a st''''') - then show ?thesis - proof (cases a) - case p1: (Pair sv tp) - with 28(1) a1 g_def gcost Some Pair s1 Method fields s2 st'_def e'_def s4 f2 n2 n3 - have "expr (CALL i xe) e\<^sub>p e cd st = Normal ((sv, tp), st'''''\stack:=stack st''', memory := memory st'''\)" and *: "gas st' \ gas (st\gas := gas st - g\)" - using expr.psimps(18)[of i xe e\<^sub>p e cd st] by (auto simp add: Let_def split: unit.split_asm) - with a1 have "gas st4' \ gas st'''''" by auto - also from 28(4)[of "()" "st\gas := gas st - g\" _ ct] g_def gcost Some Pair s1 Method fields s2 st'_def e'_def s4 f2 n2 n3 - have "\ \ gas st''''" by auto - also from 28(3)[of "()" "st\gas := gas st - g\" _ ct _ _ x1 fp _ f c x e' st' "st\gas := gas st - g\" _ st''' e'' _ cd' st'' st''' st''' "()" st''] a1 g_def gcost Some Pair s1 Method fields s2 st'_def e'_def s4 f2 n2 - have "\ \ gas st''" by auto - also have "\ \ gas st - g" - proof - - from g_def gcost have "(applyf (costs\<^sub>e (CALL i xe) e\<^sub>p e cd) \ (\g. assert Gas (\st. gas st \ g) (modify (\st. st\gas := gas st - g\)))) st = Normal ((), st\gas := gas st - g\)" by simp - moreover from e'_def have "e' = ffold_init ct (emptyEnv (address e) (sender e) (svalue e)) (fmdom ct)" by simp - moreover from st'_def have "applyf (\st. st\stack := emptyStore\) (st\gas := gas st - g\) = Normal (st', st\gas := gas st - g\)" by simp - ultimately have "\ev cda sta st'a. load False fp xe e\<^sub>p e' emptyStore st' e cd (st\gas := gas st - g\) = Normal ((ev, cda, sta), st'a) \ gas sta \ gas st' \ gas st'a \ gas (st\gas := gas st - g\) \ address ev = address e'" using 28(2)[of "()" "st\gas := gas st - g\" _ ct _ _ x1 fp "(f,c)" f c x e' st' "st\gas := gas st - g\"] using Some Pair s1 Method fields s2 by blast - thus ?thesis using st'_def s4 f2 by auto - qed - finally show ?thesis by simp - qed - next - case (e x) - with 28(1) a1 g_def gcost Some Pair s1 Method fields s2 st'_def e'_def s4 f2 n2 show ?thesis using expr.psimps(18)[of i xe e\<^sub>p e cd st] by (auto simp add:Let_def split:unit.split_asm) - qed - next - case (e x) - with 28(1) a1 g_def gcost Some Pair s1 Method fields s2 st'_def e'_def s4 f2 show ?thesis using expr.psimps(18)[of i xe e\<^sub>p e cd st] by (auto split:unit.split_asm) - qed - qed - next - case (e x) - with 28(1) a1 g_def gcost Some Pair s1 Method fields s2 st'_def e'_def show ?thesis using expr.psimps(18)[of i xe e\<^sub>p e cd st] by auto - qed - qed - qed - next - case (Var x2) - with 28(1) a1 g_def gcost Some Pair s1 show ?thesis using expr.psimps(18) by simp - qed - qed - qed - qed - qed - qed -next - case (29 ad i xe val e\<^sub>p e cd st) - define g where "g = costs\<^sub>e (ECALL ad i xe val) e\<^sub>p e cd st" - show ?case - proof (rule allI[THEN allI, THEN allI, OF impI]) - fix st4' v4 t4 assume a1: "expr (ECALL ad i xe val) e\<^sub>p e cd st = Normal ((v4, t4), st4')" - show "gas st4' \ gas st" - proof (cases) - assume "gas st \ g" - with 29 g_def a1 show ?thesis using expr.psimps by simp - next - assume gcost: "\ gas st \ g" - then show ?thesis - proof (cases "expr ad e\<^sub>p e cd (st\gas := gas st - g\)") - case (n a st') - then show ?thesis - proof (cases a) - case (Pair a b) - then show ?thesis - proof (cases a) - case (KValue adv) - then show ?thesis - proof (cases b) - case (Value x1) - then show ?thesis - proof (cases x1) - case (TSInt x1) - with 29(1) a1 g_def gcost n Pair KValue Value show ?thesis using expr.psimps(19)[of ad i xe val e\<^sub>p e cd st] by simp - next - case (TUInt x2) - with 29(1) a1 g_def gcost n Pair KValue Value show ?thesis using expr.psimps(19)[of ad i xe val e\<^sub>p e cd st] by simp - next - case TBool - with 29(1) a1 g_def gcost n Pair KValue Value show ?thesis using expr.psimps(19)[of ad i xe val e\<^sub>p e cd st] by simp - next - case TAddr - then show ?thesis - proof (cases "fmlookup e\<^sub>p adv") - case None - with 29(1) a1 g_def gcost n Pair KValue Value TAddr show ?thesis using expr.psimps(19)[of ad i xe val e\<^sub>p e cd st] by simp - next - case (Some a) - then show ?thesis - proof (cases a) - case p2: (Pair ct _) - then show ?thesis - proof (cases "fmlookup ct i") - case None - with 29(1) a1 g_def gcost n Pair KValue Value TAddr Some p2 show ?thesis using expr.psimps(19) by simp - next - case s1: (Some a) - then show ?thesis - proof (cases a) - case (Method x1) - then show ?thesis - proof (cases x1) - case (fields fp f c) - then show ?thesis - proof (cases c) - case None - with 29(1) a1 g_def gcost n Pair KValue Value TAddr Some p2 s1 Method fields show ?thesis using expr.psimps(19) by simp - next - case s2: (Some x) - then show ?thesis - proof (cases "expr val e\<^sub>p e cd st'") - case n1: (n kv st'') - then show ?thesis - proof (cases kv) - case p3: (Pair a b) - then show ?thesis - proof (cases a) - case k1: (KValue v) - then show ?thesis - proof (cases b) - case v1: (Value t) - define stl e' - where "stl = st''\stack:=emptyStore, memory:=emptyStore\" - and "e' = ffold (init ct) (emptyEnv adv (address e) v) (fmdom ct)" - then show ?thesis - proof (cases "load True fp xe e\<^sub>p e' emptyStore stl e cd st''") - case s4: (n a st''') - then show ?thesis - proof (cases a) - case f2: (fields e'' cd' st'''') - then show ?thesis - proof (cases "transfer (address e) adv v (accounts st'''')") - case n2: None - with 29(1) a1 g_def gcost n Pair KValue Value TAddr Some p2 s1 Method fields s2 n1 p3 k1 v1 stl_def e'_def s4 f2 show ?thesis using expr.psimps(19) by simp - next - case s3: (Some acc) - show ?thesis - proof (cases "stmt f e\<^sub>p e'' cd' (st''''\accounts:=acc\)") - case n2: (n a st''''') - then show ?thesis - proof (cases "expr x e\<^sub>p e'' cd' st'''''") - case n3: (n a st'''''') - then show ?thesis - proof (cases a) - case p1: (Pair sv tp) - with 29(1) a1 g_def gcost n Pair KValue Value TAddr Some p2 s1 Method fields s2 n1 p3 k1 v1 s3 stl_def e'_def s4 f2 n2 n3 - have "expr (ECALL ad i xe val) e\<^sub>p e cd st = Normal ((sv, tp), st''''''\stack:=stack st''', memory := memory st'''\)" - using expr.psimps(19)[of ad i xe val e\<^sub>p e cd st] by (auto simp add: Let_def split: unit.split_asm) - with a1 have "gas st4' \ gas st''''''" by auto - also from 29(6)[of "()" "st\gas := gas st - g\" _ st' _ _ adv _ _ ct _ _ x1 fp "(f,c)" f c x kv st'' _ b v t] a1 g_def gcost n Pair KValue Value TAddr Some p2 s1 Method fields s2 n1 p3 k1 v1 s3 stl_def e'_def s4 f2 n2 n3 - have "\ \ gas st'''''" by auto - also from 29(5)[OF _ n Pair KValue Value TAddr Some p2 s1 Method fields _ s2 n1 p3 k1 v1 _ _ s4 f2 _ _ _, of "()" f cd' st'''' st''' st''' acc] f2 s3 stl_def e'_def n2 n3 a1 g_def gcost - have "\ \ gas (st''''\accounts:=acc\)" by auto - also have "\ \ gas stl" - proof - - from g_def gcost have "(applyf (costs\<^sub>e (ECALL ad i xe val) e\<^sub>p e cd) \ (\g. assert Gas (\st. gas st \ g) (modify (\st. st\gas := gas st - g\)))) st = Normal ((), st\gas := gas st - g\)" by simp - moreover from e'_def have "e' = ffold_init ct (emptyEnv adv (address e) v) (fmdom ct)" by simp - moreover from n1 have "expr val e\<^sub>p e cd st' = Normal (kv, st'')" by simp - moreover from stl_def have "applyf (\st. st\stack := emptyStore, memory := emptyStore\) st'' = Normal (stl, st'')" by simp - moreover have "applyf accounts st'' = Normal ((accounts st''), st'')" by simp - ultimately have "\ev cda sta st'a. load True fp xe e\<^sub>p e' emptyStore stl e cd st'' = Normal ((ev, cda, sta), st'a) \ gas sta \ gas stl \ gas st'a \ gas st'' \ address ev = address e'" using 29(4)[of "()" "st\gas := gas st - g\" _ st' _ _ adv _ _ ct _ _ x1 fp "(f,c)" f c x kv st'' _ b v t] a1 g_def gcost n Pair KValue Value TAddr Some p2 s1 Method fields s2 n1 p3 k1 v1 s3 stl_def e'_def s4 f2 n2 n3 by blast - thus ?thesis using stl_def s4 f2 by auto - qed - also from stl_def have "\ \ gas st''" by simp - also from 29(3)[of "()" "st\gas := gas st - g\" _ st' _ _ adv _ _ ct _ _ x1 fp "(f,c)" f c x] a1 g_def gcost n Pair KValue Value TAddr Some p2 s1 Method fields s2 n1 p3 k1 v1 s3 stl_def e'_def s4 f2 n2 n3 - have "\ \ gas st'" by (auto split:unit.split_asm) - also from 29(2)[of "()" "st\gas := gas st - g\"] a1 g_def gcost n Pair KValue Value TAddr Some p2 s1 Method fields s2 n1 p3 k1 v1 s3 stl_def e'_def s4 f2 n2 n3 - have "\ \ gas (st\gas := gas st - g\)" by simp - finally show ?thesis by simp - qed - next - case (e x) - with 29(1) a1 g_def gcost n Pair KValue Value TAddr Some p2 s1 Method fields s2 n1 p3 k1 v1 s3 stl_def e'_def s4 f2 n2 show ?thesis using expr.psimps(19)[of ad i xe val e\<^sub>p e cd st] by simp - qed - next - case (e x) - with 29(1) a1 g_def gcost n Pair KValue Value TAddr Some p2 s1 Method fields s2 n1 p3 k1 v1 s3 stl_def e'_def s4 f2 show ?thesis using expr.psimps(19)[of ad i xe val e\<^sub>p e cd st] by simp - qed - qed - qed - next - case (e x) - with 29(1) a1 g_def gcost n Pair KValue Value TAddr Some p2 s1 Method fields s2 n1 p3 k1 v1 stl_def e'_def show ?thesis using expr.psimps(19)[of ad i xe val e\<^sub>p e cd st] by simp - qed - next - case (Calldata x2) - with 29(1) a1 g_def gcost n Pair KValue Value TAddr Some p2 s1 Method fields s2 n1 p3 k1 show ?thesis using expr.psimps(19)[of ad i xe val e\<^sub>p e cd st] by simp - next - case (Memory x3) - with 29(1) a1 g_def gcost n Pair KValue Value TAddr Some p2 s1 Method fields s2 n1 p3 k1 show ?thesis using expr.psimps(19)[of ad i xe val e\<^sub>p e cd st] by simp - next - case (Storage x4) - with 29(1) a1 g_def gcost n Pair KValue Value TAddr Some p2 s1 Method fields s2 n1 p3 k1 show ?thesis using expr.psimps(19)[of ad i xe val e\<^sub>p e cd st] by simp - qed - next - case (KCDptr x2) - with 29(1) a1 g_def gcost n Pair KValue Value TAddr Some p2 s1 Method fields s2 n1 p3 show ?thesis using expr.psimps(19)[of ad i xe val e\<^sub>p e cd st] by simp - next - case (KMemptr x3) - with 29(1) a1 g_def gcost n Pair KValue Value TAddr Some p2 s1 Method fields s2 n1 p3 show ?thesis using expr.psimps(19)[of ad i xe val e\<^sub>p e cd st] by simp - next - case (KStoptr x4) - with 29(1) a1 g_def gcost n Pair KValue Value TAddr Some p2 s1 Method fields s2 n1 p3 show ?thesis using expr.psimps(19)[of ad i xe val e\<^sub>p e cd st] by simp - qed - qed - next - case n2: (e x) - with 29(1) a1 g_def gcost n Pair KValue Value TAddr Some p2 s1 Method fields s2 show ?thesis using expr.psimps(19)[of ad i xe val e\<^sub>p e cd st] by simp - qed - qed - qed - next - case (Var x2) - with 29(1) a1 g_def gcost n Pair KValue Value TAddr Some p2 s1 show ?thesis using expr.psimps(19)[of ad i xe val e\<^sub>p e cd st] by simp - qed - qed - qed - qed - qed - next - case (Calldata x2) - with 29(1) a1 g_def gcost n Pair KValue show ?thesis using expr.psimps(19)[of ad i xe val e\<^sub>p e cd st] by simp - next - case (Memory x3) - with 29(1) a1 g_def gcost n Pair KValue show ?thesis using expr.psimps(19)[of ad i xe val e\<^sub>p e cd st] by simp - next - case (Storage x4) - with 29(1) a1 g_def gcost n Pair KValue show ?thesis using expr.psimps(19)[of ad i xe val e\<^sub>p e cd st] by simp - qed - next - case (KCDptr x2) - with 29(1) a1 g_def gcost n Pair show ?thesis using expr.psimps(19)[of ad i xe val e\<^sub>p e cd st] by simp - next - case (KMemptr x3) - with 29(1) a1 g_def gcost n Pair show ?thesis using expr.psimps(19)[of ad i xe val e\<^sub>p e cd st] by simp - next - case (KStoptr x4) - with 29(1) a1 g_def gcost n Pair show ?thesis using expr.psimps(19)[of ad i xe val e\<^sub>p e cd st] by simp - qed - qed - next - case (e _) - with 29(1) a1 g_def gcost show ?thesis using expr.psimps(19)[of ad i xe val e\<^sub>p e cd st] by simp - qed - qed - qed -next - case (30 cp i\<^sub>p t\<^sub>p pl e el e\<^sub>p e\<^sub>v' cd' st' e\<^sub>v cd st) - then show ?case - proof (cases "expr e e\<^sub>p e\<^sub>v cd st") - case (n a st'') - then show ?thesis - proof (cases a) - case (Pair v t) - then show ?thesis - proof (cases "decl i\<^sub>p t\<^sub>p (Some (v,t)) cp cd (memory st'') cd' e\<^sub>v' st'") - case n2: (n a' st''') - then show ?thesis - proof (cases a') - case f2: (Pair cd'' e\<^sub>v'') - show ?thesis - proof (rule allI[THEN allI, THEN allI, THEN allI, OF impI]) - fix ev xa xaa xaaa assume load_def: "load cp ((i\<^sub>p, t\<^sub>p) # pl) (e # el) e\<^sub>p e\<^sub>v' cd' st' e\<^sub>v cd st = Normal ((ev, xa, xaa), xaaa)" - with 30(1) n Pair n2 f2 have "load cp ((i\<^sub>p, t\<^sub>p) # pl) (e # el) e\<^sub>p e\<^sub>v' cd' st' e\<^sub>v cd st = load cp pl el e\<^sub>p e\<^sub>v'' cd'' st''' e\<^sub>v cd st''" using load.psimps(1)[of cp i\<^sub>p t\<^sub>p pl e el e\<^sub>p e\<^sub>v' cd' st' e\<^sub>v cd st] by simp - with load_def have "load cp pl el e\<^sub>p e\<^sub>v'' cd'' st''' e\<^sub>v cd st'' = Normal ((ev, xa, xaa), xaaa)" by simp - with n Pair n2 f2 have "gas xaa \ gas st''' \ gas xaaa \ gas st'' \ address ev = address e\<^sub>v''" using 30(3)[of a st'' v t st'' st'' "()" st' a' st''' cd'' e\<^sub>v'' st''' st''' "()" st''] by simp - moreover from n Pair have "gas st'' \ gas st" using 30(2) by simp - moreover from n2 f2 have " address e\<^sub>v'' = address e\<^sub>v'" and "gas st''' \ gas st'" using decl_gas_address by auto - ultimately show "gas xaa \ gas st' \ gas xaaa \ gas st \ address ev = address e\<^sub>v'" by simp - qed - qed - next - case (e x) - with 30(1) n Pair show ?thesis using load.psimps(1) by simp - qed - qed - next - case (e x) - with 30(1) show ?thesis using load.psimps(1) by simp - qed -next - case (31 we wf wg wh wi wj wk st) - then show ?case using load.psimps(2) by auto -next - case (32 wl wm wn wo wp wq wr st) - then show ?case using load.psimps(3)[of wl wm wn wo wp wq wr] by auto -next - case (33 ws wt wu wv cd e\<^sub>v s st) - then show ?case using load.psimps(4)[of ws wt wu wv cd e\<^sub>v s st] by auto -next - case (34 i e\<^sub>p e cd st) - show ?case - proof (rule allI[THEN allI, THEN allI, OF impI]) - fix st3' xa xaa assume "rexp (L.Id i) e\<^sub>p e cd st = Normal ((st3', xa), xaa)" - then show "gas xaa \ gas st" using 34(1) rexp.psimps(1) by (simp split: option.split_asm Denvalue.split_asm Stackvalue.split_asm prod.split_asm Type.split_asm STypes.split_asm) - qed -next - case (35 i r e\<^sub>p e cd st) - show ?case - proof (rule allI[THEN allI, THEN allI, OF impI]) - fix st3' xa xaa assume rexp_def: "rexp (Ref i r) e\<^sub>p e cd st = Normal ((st3', xa), xaa)" - show "gas xaa \ gas st" - proof (cases "fmlookup (denvalue e) i") - case None - with 35(1) show ?thesis using rexp.psimps rexp_def by simp - next - case (Some a) - then show ?thesis - proof (cases a) - case (Pair tp b) - then show ?thesis - proof (cases b) - case (Stackloc l) - then show ?thesis - proof (cases "accessStore l (stack st)") - case None - with 35(1) Some Pair Stackloc show ?thesis using rexp.psimps(2) rexp_def by simp - next - case s1: (Some a) - then show ?thesis - proof (cases a) - case (KValue x1) - with 35(1) Some Pair Stackloc s1 show ?thesis using rexp.psimps(2) rexp_def by simp - next - case (KCDptr l') - with 35 Some Pair Stackloc s1 show ?thesis using rexp.psimps(2)[of i r e\<^sub>p e cd st] rexp_def by (simp split: option.split_asm Memoryvalue.split_asm MTypes.split_asm prod.split_asm Type.split_asm StateMonad.result.split_asm) - next - case (KMemptr x3) - with 35 Some Pair Stackloc s1 show ?thesis using rexp.psimps(2)[of i r e\<^sub>p e cd st] rexp_def by (simp split: option.split_asm Memoryvalue.split_asm MTypes.split_asm prod.split_asm Type.split_asm StateMonad.result.split_asm) - next - case (KStoptr x4) - with 35 Some Pair Stackloc s1 show ?thesis using rexp.psimps(2)[of i r e\<^sub>p e cd st] rexp_def by (simp split: option.split_asm STypes.split_asm prod.split_asm Type.split_asm StateMonad.result.split_asm) - qed - qed - next - case (Storeloc x2) - with 35 Some Pair show ?thesis using rexp.psimps rexp_def by (simp split: option.split_asm STypes.split_asm prod.split_asm Type.split_asm StateMonad.result.split_asm) - qed - qed - qed - qed -next - case (36 e\<^sub>p e cd st) - then show ?case using stmt.psimps(1) by simp -next - case (37 lv ex e\<^sub>p env cd st) - define g where "g = costs (ASSIGN lv ex) e\<^sub>p env cd st" - show ?case - proof (rule allI[OF impI]) - fix st6' - assume stmt_def: "stmt (ASSIGN lv ex) e\<^sub>p env cd st = Normal ((), st6')" - then show "gas st6' \ gas st" - proof cases - assume "gas st \ g" - with 37 stmt_def show ?thesis using stmt.psimps(2) g_def by simp - next - assume "\ gas st \ g" - show ?thesis - proof (cases "expr ex e\<^sub>p env cd (st\gas := gas st - g\)") - case (n a st') - then show ?thesis - proof (cases a) - case (Pair b c) - then show ?thesis - proof (cases b) - case (KValue v) - then show ?thesis - proof (cases c) - case (Value t) - then show ?thesis - proof (cases "lexp lv e\<^sub>p env cd st'") - case n2: (n a st'') - then show ?thesis - proof (cases a) - case p1: (Pair a b) - then show ?thesis - proof (cases a) - case (LStackloc l) - then show ?thesis - proof (cases b) - case v2: (Value t') - then show ?thesis - proof (cases "convert t t' v ") - case None - with 37(1) stmt_def `\ gas st \ g` n Pair KValue Value n2 p1 LStackloc v2 show ?thesis using stmt.psimps(2) g_def by simp - next - case s3: (Some a) - then show ?thesis - proof (cases a) - case p2: (Pair v' b) - with 37(1) `\ gas st \ g` n Pair KValue Value n2 p1 LStackloc v2 s3 - have "stmt (ASSIGN lv ex) e\<^sub>p env cd st = Normal ((), st'' \stack := updateStore l (KValue v') (stack st'')\)" - using stmt.psimps(2) g_def by simp - with stmt_def have "st6'= (st''\stack := updateStore l (KValue v') (stack st'')\)" by simp - moreover from 37(3) `\ gas st \ g` n Pair KValue Value n2 p1 have "gas st'' \ gas st'" using g_def by simp - moreover from 37(2) `\ gas st \ g` n Pair KValue Value n2 p2 have "gas st' \ gas st" using g_def by simp - ultimately show ?thesis by simp - qed - qed - next - case (Calldata x2) - with 37(1) stmt_def `\ gas st \ g` n Pair KValue Value n2 p1 LStackloc show ?thesis using stmt.psimps(2) g_def by simp - next - case (Memory x3) - with 37(1) stmt_def `\ gas st \ g` n Pair KValue Value n2 p1 LStackloc show ?thesis using stmt.psimps(2) g_def by simp - next - case (Storage x4) - with 37(1) stmt_def `\ gas st \ g` n Pair KValue Value n2 p1 LStackloc show ?thesis using stmt.psimps(2) g_def by simp - qed - next - case (LMemloc l) - then show ?thesis - proof (cases b) - case v2: (Value t') - with 37(1) stmt_def `\ gas st \ g` n Pair KValue Value n2 p1 LMemloc show ?thesis using stmt.psimps(2) g_def by simp - next - case (Calldata x2) - with 37(1) stmt_def `\ gas st \ g` n Pair KValue Value n2 p1 LMemloc show ?thesis using stmt.psimps(2) g_def by simp - next - case (Memory x3) - then show ?thesis - proof (cases x3) - case (MTArray x11 x12) - with 37(1) stmt_def `\ gas st \ g` n Pair KValue Value n2 p1 LMemloc Memory show ?thesis using stmt.psimps(2) g_def by simp - next - case (MTValue t') - then show ?thesis - proof (cases "convert t t' v ") - case None - with 37(1) stmt_def `\ gas st \ g` n Pair KValue Value n2 p1 LMemloc Memory MTValue show ?thesis using stmt.psimps(2) g_def by simp - next - case s3: (Some a) - then show ?thesis - proof (cases a) - case p2: (Pair v' b) - with 37(1) `\ gas st \ g` n Pair KValue Value n2 p1 LMemloc Memory MTValue s3 - have "stmt (ASSIGN lv ex) e\<^sub>p env cd st = Normal ((), st'' \memory := updateStore l (MValue v') (memory st'')\)" - using stmt.psimps(2) g_def by simp - with stmt_def have "st6'= (st''\memory := updateStore l (MValue v') (memory st'')\)" by simp - moreover from 37(3) `\ gas st \ g` n Pair KValue Value n2 p1 have "gas st'' \ gas st'" using g_def by simp - moreover from 37(2) `\ gas st \ g` n Pair KValue Value n2 p1 have "gas st' \ gas st" using g_def by simp - ultimately show ?thesis by simp - qed - qed - qed - next - case (Storage x4) - with 37(1) stmt_def `\ gas st \ g` n Pair KValue Value n2 p1 LMemloc show ?thesis using stmt.psimps(2) g_def by simp - qed - next - case (LStoreloc l) - then show ?thesis - proof (cases b) - case v2: (Value t') - with 37(1) stmt_def `\ gas st \ g` n Pair KValue Value n2 p1 LStoreloc show ?thesis using stmt.psimps(2) g_def by simp - next - case (Calldata x2) - with 37(1) stmt_def `\ gas st \ g` n Pair KValue Value n2 p1 LStoreloc show ?thesis using stmt.psimps(2) g_def by simp - next - case (Memory x3) - with 37(1) stmt_def `\ gas st \ g` n Pair KValue Value n2 p1 LStoreloc show ?thesis using stmt.psimps(2) g_def by simp - next - case (Storage x4) - then show ?thesis - proof (cases x4) - case (STArray x11 x12) - with 37(1) stmt_def `\ gas st \ g` n Pair KValue Value n2 p1 LStoreloc Storage show ?thesis using stmt.psimps(2) g_def by simp - next - case (STMap x21 x22) - with 37(1) stmt_def `\ gas st \ g` n Pair KValue Value n2 p1 LStoreloc Storage show ?thesis using stmt.psimps(2) g_def by simp - next - case (STValue t') - then show ?thesis - proof (cases "convert t t' v ") - case None - with 37(1) stmt_def `\ gas st \ g` n Pair KValue Value n2 p1 LStoreloc Storage STValue show ?thesis using stmt.psimps(2) g_def by simp - next - case s3: (Some a) - then show ?thesis - proof (cases a) - case p2: (Pair v' b) - then show ?thesis - proof (cases "fmlookup (storage st'') (address env)") - case None - with 37(1) stmt_def `\ gas st \ g` n Pair KValue Value n2 p1 LStoreloc Storage STValue s3 p2 show ?thesis using stmt.psimps(2) g_def by simp - next - case s4: (Some s) - with 37(1) `\ gas st \ g` n Pair KValue Value n2 p1 LStoreloc Storage STValue s3 p2 - have "stmt (ASSIGN lv ex) e\<^sub>p env cd st = Normal ((), st'' \storage := fmupd (address env) (fmupd l v' s) (storage st'')\)" - using stmt.psimps(2) g_def by simp - with stmt_def have "st6'= st'' \storage := fmupd (address env) (fmupd l v' s) (storage st'')\" by simp - moreover from 37(3) `\ gas st \ g` n Pair KValue Value n2 p1 have "gas st'' \ gas st'" using g_def by simp - moreover from 37(2) `\ gas st \ g` n Pair KValue Value n2 p1 have "gas st' \ gas st" using g_def by simp - ultimately show ?thesis by simp - qed - qed - qed - qed - qed - qed - qed - next - case (e x) - with 37(1) stmt_def `\ gas st \ g` n Pair KValue Value show ?thesis using stmt.psimps(2) g_def by simp - qed - next - case (Calldata x2) - with 37(1) stmt_def `\ gas st \ g` n Pair KValue show ?thesis using stmt.psimps(2) g_def by simp - next - case (Memory x3) - with 37(1) stmt_def `\ gas st \ g` n Pair KValue show ?thesis using stmt.psimps(2) g_def by simp - next - case (Storage x4) - with 37(1) stmt_def `\ gas st \ g` n Pair KValue show ?thesis using stmt.psimps(2) g_def by simp - qed - next - case (KCDptr p) - then show ?thesis - proof (cases c) - case (Value x1) - with 37(1) stmt_def `\ gas st \ g` n Pair KCDptr show ?thesis using stmt.psimps(2) g_def by simp - next - case (Calldata x2) - then show ?thesis - proof (cases x2) - case (MTArray x t) - then show ?thesis - proof (cases "lexp lv e\<^sub>p env cd st'") - case n2: (n a st'') - then show ?thesis - proof (cases a) - case p2: (Pair a b) - then show ?thesis - proof (cases a) - case (LStackloc l) - then show ?thesis - proof (cases b) - case v2: (Value t') - with 37(1) stmt_def `\ gas st \ g` n Pair KCDptr Calldata MTArray n2 p2 LStackloc show ?thesis using stmt.psimps(2) g_def by simp - next - case c2: (Calldata x2) - with 37(1) stmt_def `\ gas st \ g` n Pair KCDptr Calldata MTArray n2 p2 LStackloc show ?thesis using stmt.psimps(2) g_def by simp - next - case (Memory x3) - with 37(1) `\ gas st \ g` n Pair KCDptr Calldata MTArray n2 p2 LStackloc - have "stmt (ASSIGN lv ex) e\<^sub>p env cd st = Normal ((), st'' \stack := updateStore l (KCDptr p) (stack st'')\)" - using stmt.psimps(2) g_def by simp - with stmt_def have "st6'= (st''\stack := updateStore l (KCDptr p) (stack st'')\)" by simp - moreover from 37(4) `\ gas st \ g` n Pair KCDptr Calldata MTArray n2 p2 have "gas st'' \ gas st'" using g_def by simp - moreover from 37(2) `\ gas st \ g` n Pair have "gas st' \ gas st" using g_def by simp - ultimately show ?thesis by simp - next - case (Storage x4) - then show ?thesis - proof (cases "accessStore l (stack st'')") - case None - with 37(1) stmt_def `\ gas st \ g` n Pair KCDptr Calldata MTArray n2 p2 LStackloc Storage show ?thesis using stmt.psimps(2) g_def by simp - next - case s3: (Some a) - then show ?thesis - proof (cases a) - case (KValue x1) - with 37(1) stmt_def `\ gas st \ g` n Pair KCDptr Calldata MTArray n2 p2 LStackloc Storage s3 show ?thesis using stmt.psimps(2) g_def by simp - next - case c3: (KCDptr x2) - with 37(1) stmt_def `\ gas st \ g` n Pair KCDptr Calldata MTArray n2 p2 LStackloc Storage s3 show ?thesis using stmt.psimps(2) g_def by simp - next - case (KMemptr x3) - with 37(1) stmt_def `\ gas st \ g` n Pair KCDptr Calldata MTArray n2 p2 LStackloc Storage s3 show ?thesis using stmt.psimps(2) g_def by simp - next - case (KStoptr p') - then show ?thesis - proof (cases "fmlookup (storage st'') (address env)") - case None - with 37(1) stmt_def `\ gas st \ g` n Pair KCDptr Calldata MTArray n2 p2 LStackloc Storage s3 KStoptr show ?thesis using stmt.psimps(2) g_def by simp - next - case s4: (Some s) - then show ?thesis - proof (cases "cpm2s p p' x t cd s") - case None - with 37(1) stmt_def `\ gas st \ g` n Pair KCDptr Calldata MTArray n2 p2 LStackloc Storage s3 KStoptr s4 show ?thesis using stmt.psimps(2) g_def by simp - next - case (Some s') - with 37(1) `\ gas st \ g` n Pair KCDptr Calldata MTArray n2 p2 LStackloc Storage s3 KStoptr s4 - have "stmt (ASSIGN lv ex) e\<^sub>p env cd st = Normal ((), st'' \storage := fmupd (address env) s' (storage st'')\)" - using stmt.psimps(2) g_def by simp - with stmt_def have "st6'= st'' \storage := fmupd (address env) s' (storage st'')\" by simp - moreover from 37(4) `\ gas st \ g` n Pair KCDptr Calldata MTArray n2 p2 have "gas st'' \ gas st'" using g_def by simp - moreover from 37(2) `\ gas st \ g` n Pair have "gas st' \ gas st" using g_def by simp - ultimately show ?thesis by simp - qed - qed - qed - qed - qed - next - case (LMemloc l) - then show ?thesis - proof (cases "cpm2m p l x t cd (memory st'')") - case None - with 37(1) stmt_def `\ gas st \ g` n Pair KCDptr Calldata MTArray n2 p2 LMemloc show ?thesis using stmt.psimps(2) g_def by simp - next - case (Some m) - with 37(1) `\ gas st \ g` n Pair KCDptr Calldata MTArray n2 p2 LMemloc - have "stmt (ASSIGN lv ex) e\<^sub>p env cd st = Normal ((), st'' \memory := m\)" - using stmt.psimps(2) g_def by simp - with stmt_def have "st6'= (st''\memory := m\)" by simp - moreover from 37(4) `\ gas st \ g` n Pair KCDptr Calldata MTArray n2 p2 have "gas st'' \ gas st'" using g_def by simp - moreover from 37(2) `\ gas st \ g` n Pair have "gas st' \ gas st" using g_def by simp - ultimately show ?thesis by simp - qed - next - case (LStoreloc l) - then show ?thesis - proof (cases "fmlookup (storage st'') (address env)") - case None - with 37(1) stmt_def `\ gas st \ g` n Pair KCDptr Calldata MTArray n2 p2 LStoreloc show ?thesis using stmt.psimps(2) g_def by simp - next - case s4: (Some s) - then show ?thesis - proof (cases "cpm2s p l x t cd s") - case None - with 37(1) stmt_def `\ gas st \ g` n Pair KCDptr Calldata MTArray n2 p2 LStoreloc s4 show ?thesis using stmt.psimps(2) g_def by simp - next - case (Some s') - with 37(1) `\ gas st \ g` n Pair KCDptr Calldata MTArray n2 p2 LStoreloc s4 - have "stmt (ASSIGN lv ex) e\<^sub>p env cd st = Normal ((), st'' \storage := fmupd (address env) s' (storage st'')\)" - using stmt.psimps(2) g_def by simp - with stmt_def have "st6'= (st''\storage := fmupd (address env) s' (storage st'')\)" by simp - moreover from 37(4) `\ gas st \ g` n Pair KCDptr Calldata MTArray n2 p2 have "gas st'' \ gas st'" using g_def by simp - moreover from 37(2) `\ gas st \ g` n Pair have "gas st' \ gas st" using g_def by simp - ultimately show ?thesis by simp - qed - qed - qed - qed - next - case (e x) - with 37(1) stmt_def `\ gas st \ g` n Pair KCDptr Calldata MTArray show ?thesis using stmt.psimps(2) g_def by simp - qed - next - case (MTValue x2) - with 37(1) stmt_def `\ gas st \ g` n Pair KCDptr Calldata show ?thesis using stmt.psimps(2) g_def by simp - qed - next - case (Memory x3) - with 37(1) stmt_def `\ gas st \ g` n Pair KCDptr show ?thesis using stmt.psimps(2) g_def by simp - next - case (Storage x4) - with 37(1) stmt_def `\ gas st \ g` n Pair KCDptr show ?thesis using stmt.psimps(2) g_def by simp - qed - next - case (KMemptr p) - then show ?thesis - proof (cases c) - case (Value x1) - with 37(1) stmt_def `\ gas st \ g` n Pair KMemptr show ?thesis using stmt.psimps(2) g_def by simp - next - case (Calldata x2) - with 37(1) stmt_def `\ gas st \ g` n Pair KMemptr show ?thesis using stmt.psimps(2) g_def by simp - next - case (Memory x3) - then show ?thesis - proof (cases x3) - case (MTArray x t) - then show ?thesis - proof (cases "lexp lv e\<^sub>p env cd st'") - case n2: (n a st'') - then show ?thesis - proof (cases a) - case p2: (Pair a b) - then show ?thesis - proof (cases a) - case (LStackloc l) - then show ?thesis - proof (cases b) - case v2: (Value t') - with 37(1) stmt_def `\ gas st \ g` n Pair KMemptr Memory MTArray n2 p2 LStackloc show ?thesis using stmt.psimps(2) g_def by simp - next - case c2: (Calldata x2) - with 37(1) stmt_def `\ gas st \ g` n Pair KMemptr Memory MTArray n2 p2 LStackloc show ?thesis using stmt.psimps(2) g_def by simp - next - case m2: (Memory x3) - with 37(1) `\ gas st \ g` n Pair KMemptr Memory MTArray n2 p2 LStackloc - have "stmt (ASSIGN lv ex) e\<^sub>p env cd st = Normal ((), st'' \stack := updateStore l (KMemptr p) (stack st'')\)" - using stmt.psimps(2) g_def by simp - with stmt_def have "st6'= (st''\stack := updateStore l (KMemptr p) (stack st'')\)" by simp - moreover from 37(5) `\ gas st \ g` n Pair KMemptr Memory MTArray n2 p2 have "gas st'' \ gas st'" using g_def by simp - moreover from 37(2) `\ gas st \ g` n Pair have "gas st' \ gas st" using g_def by simp - ultimately show ?thesis by simp - next - case (Storage x4) - then show ?thesis - proof (cases "accessStore l (stack st'')") - case None - with 37(1) stmt_def `\ gas st \ g` n Pair KMemptr Memory MTArray n2 p2 LStackloc Storage show ?thesis using stmt.psimps(2) g_def by simp - next - case s3: (Some a) - then show ?thesis - proof (cases a) - case (KValue x1) - with 37(1) stmt_def `\ gas st \ g` n Pair KMemptr Memory MTArray n2 p2 LStackloc Storage s3 show ?thesis using stmt.psimps(2) g_def by simp - next - case c3: (KCDptr x2) - with 37(1) stmt_def `\ gas st \ g` n Pair KMemptr Memory MTArray n2 p2 LStackloc Storage s3 show ?thesis using stmt.psimps(2) g_def by simp - next - case m3: (KMemptr x3) - with 37(1) stmt_def `\ gas st \ g` n Pair KMemptr Memory MTArray n2 p2 LStackloc Storage s3 show ?thesis using stmt.psimps(2) g_def by simp - next - case (KStoptr p') - then show ?thesis - proof (cases "fmlookup (storage st'') (address env)") - case None - with 37(1) stmt_def `\ gas st \ g` n Pair KMemptr Memory MTArray n2 p2 LStackloc Storage s3 KStoptr show ?thesis using stmt.psimps(2) g_def by simp - next - case s4: (Some s) - then show ?thesis - proof (cases "cpm2s p p' x t (memory st'') s") - case None - with 37(1) stmt_def `\ gas st \ g` n Pair KMemptr Memory MTArray n2 p2 LStackloc Storage s3 KStoptr s4 show ?thesis using stmt.psimps(2) g_def by simp - next - case (Some s') - with 37(1) `\ gas st \ g` n Pair KMemptr Memory MTArray n2 p2 LStackloc Storage s3 KStoptr s4 - have "stmt (ASSIGN lv ex) e\<^sub>p env cd st = Normal ((), st'' \storage := fmupd (address env) s' (storage st'')\)" - using stmt.psimps(2) g_def by simp - with stmt_def have "st6'= (st''\storage := fmupd (address env) s' (storage st'')\)" by simp - moreover from 37(5) `\ gas st \ g` n Pair KMemptr Memory MTArray n2 p2 have "gas st'' \ gas st'" using g_def by simp - moreover from 37(2) `\ gas st \ g` n Pair have "gas st' \ gas st" using g_def by simp - ultimately show ?thesis by simp - qed - qed - qed - qed - qed - next - case (LMemloc l) - with 37(1) `\ gas st \ g` n Pair KMemptr Memory MTArray n2 p2 LMemloc - have "stmt (ASSIGN lv ex) e\<^sub>p env cd st = Normal ((), st'' \memory := updateStore l (MPointer p) (memory st'')\)" - using stmt.psimps(2) g_def by simp - with stmt_def have "st6'= st'' \memory := updateStore l (MPointer p) (memory st'')\" by simp - moreover from 37(5) `\ gas st \ g` n Pair KMemptr Memory MTArray n2 p2 have "gas st'' \ gas st'" using g_def by simp - moreover from 37(2) `\ gas st \ g` n Pair have "gas st' \ gas st" using g_def by simp - ultimately show ?thesis by simp - next - case (LStoreloc l) - then show ?thesis - proof (cases "fmlookup (storage st'') (address env)") - case None - with 37(1) stmt_def `\ gas st \ g` n Pair KMemptr Memory MTArray n2 p2 LStoreloc show ?thesis using stmt.psimps(2) g_def by simp - next - case s3: (Some s) - then show ?thesis - proof (cases "cpm2s p l x t (memory st'') s") - case None - with 37(1) stmt_def `\ gas st \ g` n Pair KMemptr Memory MTArray n2 p2 LStoreloc s3 show ?thesis using stmt.psimps(2) g_def by simp - next - case (Some s') - with 37(1) `\ gas st \ g` n Pair KMemptr Memory MTArray n2 p2 LStoreloc s3 - have "stmt (ASSIGN lv ex) e\<^sub>p env cd st = Normal ((), st'' \storage := fmupd (address env) s' (storage st'')\)" - using stmt.psimps(2) g_def by simp - with stmt_def have "st6'= st''\storage := fmupd (address env) s' (storage st'')\" by simp - moreover from 37(5) `\ gas st \ g` n Pair KMemptr Memory MTArray n2 p2 have "gas st'' \ gas st'" using g_def by simp - moreover from 37(2) `\ gas st \ g` n Pair have "gas st' \ gas st" using g_def by simp - ultimately show ?thesis by simp - qed - qed - qed - qed - next - case (e x) - with 37(1) stmt_def `\ gas st \ g` n Pair KMemptr Memory MTArray show ?thesis using stmt.psimps(2) g_def by simp - qed - next - case (MTValue x2) - with 37(1) stmt_def `\ gas st \ g` n Pair KMemptr Memory show ?thesis using stmt.psimps(2) g_def by simp - qed - next - case (Storage x4) - with 37(1) stmt_def `\ gas st \ g` n Pair KMemptr show ?thesis using stmt.psimps(2) g_def by simp - qed - next - case (KStoptr p) - then show ?thesis - proof (cases c) - case (Value x1) - with 37(1) stmt_def `\ gas st \ g` n Pair KStoptr show ?thesis using stmt.psimps(2) g_def by simp - next - case (Calldata x2) - with 37(1) stmt_def `\ gas st \ g` n Pair KStoptr show ?thesis using stmt.psimps(2) g_def by simp - next - case (Memory x3) - with 37(1) stmt_def `\ gas st \ g` n Pair KStoptr show ?thesis using stmt.psimps(2) g_def by simp - next - case (Storage x4) - then show ?thesis - proof (cases x4) - case (STArray x t) - then show ?thesis - proof (cases "lexp lv e\<^sub>p env cd st'") - case n2: (n a st'') - then show ?thesis - proof (cases a) - case p2: (Pair a b) - then show ?thesis - proof (cases a) - case (LStackloc l) - then show ?thesis - proof (cases b) - case v2: (Value t') - with 37(1) stmt_def `\ gas st \ g` n Pair KStoptr Storage STArray n2 p2 LStackloc show ?thesis using stmt.psimps(2) g_def by simp - next - case c2: (Calldata x2) - with 37(1) stmt_def `\ gas st \ g` n Pair KStoptr Storage STArray n2 p2 LStackloc show ?thesis using stmt.psimps(2) g_def by simp - next - case (Memory x3) - then show ?thesis - proof (cases "accessStore l (stack st'')") - case None - with 37(1) stmt_def `\ gas st \ g` n Pair KStoptr Storage STArray n2 p2 LStackloc Memory show ?thesis using stmt.psimps(2) g_def by simp - next - case s3: (Some a) - then show ?thesis - proof (cases a) - case (KValue x1) - with 37(1) stmt_def `\ gas st \ g` n Pair KStoptr Storage STArray n2 p2 LStackloc Memory s3 show ?thesis using stmt.psimps(2) g_def by simp - next - case c3: (KCDptr x2) - with 37(1) stmt_def `\ gas st \ g` n Pair KStoptr Storage STArray n2 p2 LStackloc Memory s3 show ?thesis using stmt.psimps(2) g_def by simp - next - case (KMemptr p') - then show ?thesis - proof (cases "fmlookup (storage st'') (address env)") - case None - with 37(1) stmt_def `\ gas st \ g` n Pair KStoptr Storage STArray n2 p2 LStackloc Memory s3 KMemptr show ?thesis using stmt.psimps(2) g_def by simp - next - case s4: (Some s) - then show ?thesis - proof (cases "cps2m p p' x t s (memory st'')") - case None - with 37(1) stmt_def `\ gas st \ g` n Pair KStoptr Storage STArray n2 p2 LStackloc Memory s3 KMemptr s4 show ?thesis using stmt.psimps(2) g_def by simp - next - case (Some m) - with 37(1) `\ gas st \ g` n Pair KStoptr Storage STArray n2 p2 LStackloc Memory s3 KMemptr s4 - have "stmt (ASSIGN lv ex) e\<^sub>p env cd st = Normal ((), st'' \memory := m\)" - using stmt.psimps(2) g_def by simp - with stmt_def have "st6'= (st''\memory := m\)" by simp - moreover from 37(6) `\ gas st \ g` n Pair KStoptr Storage STArray n2 p2 have "gas st'' \ gas st'" using g_def by simp - moreover from 37(2) `\ gas st \ g` n Pair have "gas st' \ gas st" using g_def by simp - ultimately show ?thesis by simp - qed - qed - next - case sp2: (KStoptr p') - with 37(1) stmt_def `\ gas st \ g` n Pair KStoptr Storage STArray n2 p2 LStackloc Memory s3 show ?thesis using stmt.psimps(2) g_def by simp - qed - qed - next - case st2: (Storage x4) - with 37(1) `\ gas st \ g` n Pair KStoptr Storage STArray n2 p2 LStackloc - have "stmt (ASSIGN lv ex) e\<^sub>p env cd st = Normal ((), st'' \stack := updateStore l (KStoptr p) (stack st'')\)" - using stmt.psimps(2) g_def by simp - with stmt_def have "st6'= (st''\stack := updateStore l (KStoptr p) (stack st'')\)" by simp - moreover from 37(6) `\ gas st \ g` n Pair KStoptr Storage STArray n2 p2 have "gas st'' \ gas st'" using g_def by simp - moreover from 37(2) `\ gas st \ g` n Pair have "gas st' \ gas st" using g_def by simp - ultimately show ?thesis by simp - qed - next - case (LMemloc l) - then show ?thesis - proof (cases "fmlookup (storage st'') (address env)") - case None - with 37(1) stmt_def `\ gas st \ g` n Pair KStoptr Storage STArray n2 p2 LMemloc show ?thesis using stmt.psimps(2) g_def by simp - next - case s4: (Some s) - then show ?thesis - proof (cases "cps2m p l x t s (memory st'')") - case None - with 37(1) stmt_def `\ gas st \ g` n Pair KStoptr Storage STArray n2 p2 LMemloc s4 show ?thesis using stmt.psimps(2) g_def by simp - next - case (Some m) - with 37(1) `\ gas st \ g` n Pair KStoptr Storage STArray n2 p2 LMemloc s4 - have "stmt (ASSIGN lv ex) e\<^sub>p env cd st = Normal ((), st'' \memory := m\)" - using stmt.psimps(2) g_def by simp - with stmt_def have "st6'= (st''\memory := m\)" by simp - moreover from 37(6) `\ gas st \ g` n Pair KStoptr Storage STArray n2 p2 have "gas st'' \ gas st'" using g_def by simp - moreover from 37(2) `\ gas st \ g` n Pair have "gas st' \ gas st" using g_def by simp - ultimately show ?thesis by simp - qed - qed - next - case (LStoreloc l) - then show ?thesis - proof (cases "fmlookup (storage st'') (address env)") - case None - with 37(1) stmt_def `\ gas st \ g` n Pair KStoptr Storage STArray n2 p2 LStoreloc show ?thesis using stmt.psimps(2) g_def by simp - next - case s4: (Some s) - then show ?thesis - proof (cases "copy p l x t s") - case None - with 37(1) stmt_def `\ gas st \ g` n Pair KStoptr Storage STArray n2 p2 LStoreloc s4 show ?thesis using stmt.psimps(2) g_def by simp - next - case (Some s') - with 37(1) `\ gas st \ g` n Pair KStoptr Storage STArray n2 p2 LStoreloc s4 - have "stmt (ASSIGN lv ex) e\<^sub>p env cd st = Normal ((), st'' \storage := fmupd (address env) s' (storage st'')\)" - using stmt.psimps(2) g_def by simp - with stmt_def have "st6'= st''\storage := fmupd (address env) s' (storage st'')\" by simp - moreover from 37(6) `\ gas st \ g` n Pair KStoptr Storage STArray n2 p2 have "gas st'' \ gas st'" using g_def by simp - moreover from 37(2) `\ gas st \ g` n Pair have "gas st' \ gas st" using g_def by simp - ultimately show ?thesis by simp - qed - qed - qed - qed - next - case (e x) - with 37(1) stmt_def `\ gas st \ g` n Pair KStoptr Storage STArray show ?thesis using stmt.psimps(2) g_def by simp - qed - next - case (STMap t t') - then show ?thesis - proof (cases "lexp lv e\<^sub>p env cd st'") - case n2: (n a st'') - then show ?thesis - proof (cases a) - case p2: (Pair a b) - then show ?thesis - proof (cases a) - case (LStackloc l) - with 37(1) `\ gas st \ g` n Pair KStoptr Storage STMap n2 p2 - have "stmt (ASSIGN lv ex) e\<^sub>p env cd st = Normal ((), st'' \stack := updateStore l (KStoptr p) (stack st'')\)" - using stmt.psimps(2) g_def by simp - with stmt_def have "st6'= st''\stack := updateStore l (KStoptr p) (stack st'')\" by simp - moreover from 37(7) `\ gas st \ g` n Pair KStoptr Storage STMap n2 p2 have "gas st'' \ gas st'" using g_def by simp - moreover from 37(2) `\ gas st \ g` n Pair have "gas st' \ gas st" using g_def by simp - ultimately show ?thesis by simp - next - case (LMemloc x2) - with 37(1) stmt_def `\ gas st \ g` n Pair KStoptr Storage STMap n2 p2 show ?thesis using stmt.psimps(2) g_def by simp - next - case (LStoreloc x3) - with 37(1) stmt_def `\ gas st \ g` n Pair KStoptr Storage STMap n2 p2 show ?thesis using stmt.psimps(2) g_def by simp - qed - qed - next - case (e x) - with 37(1) stmt_def `\ gas st \ g` n Pair KStoptr Storage STMap show ?thesis using stmt.psimps(2) g_def by simp - qed - next - case (STValue x3) - with 37(1) stmt_def `\ gas st \ g` n Pair KStoptr Storage show ?thesis using stmt.psimps(2) g_def by simp - qed - qed - qed - qed - next - case (e x) - with 37(1) stmt_def `\ gas st \ g` show ?thesis using stmt.psimps(2) g_def by (simp split: Ex.split_asm) - qed - qed - qed -next - case (38 s1 s2 e\<^sub>p e cd st) - define g where "g = costs (COMP s1 s2) e\<^sub>p e cd st" - show ?case - proof (rule allI[OF impI]) - fix st6' - assume stmt_def: "stmt (COMP s1 s2) e\<^sub>p e cd st = Normal ((), st6')" - then show "gas st6' \ gas st" - proof cases - assume "gas st \ g" - with 38 stmt_def g_def show ?thesis using stmt.psimps(3) by simp - next - assume "\ gas st \ g" - show ?thesis - proof (cases "stmt s1 e\<^sub>p e cd (st\gas := gas st - g\)") - case (n a st') - with 38(1) stmt_def `\ gas st \ g` have "stmt (COMP s1 s2) e\<^sub>p e cd st = stmt s2 e\<^sub>p e cd st'" using stmt.psimps(3)[of s1 s2 e\<^sub>p e cd st] g_def by (simp add:Let_def split:unit.split_asm) - with 38(3)[of _ "(st\gas := gas st - g\)" _ st'] stmt_def \\ gas st \ g\ n have "gas st6' \ gas st'" using g_def by fastforce - moreover from 38(2) \\ gas st \ g\ n have "gas st' \ gas st" using g_def by simp - ultimately show ?thesis by simp - next - case (e x) - with 38 stmt_def `\ gas st \ g` show ?thesis using stmt.psimps(3)[of s1 s2 e\<^sub>p e cd st] g_def by (simp split: Ex.split_asm) - qed - qed - qed -next - case (39 ex s1 s2 e\<^sub>p e cd st) - define g where "g = costs (ITE ex s1 s2) e\<^sub>p e cd st" - show ?case - proof (rule allI[OF impI]) - fix st6' - assume stmt_def: "stmt (ITE ex s1 s2) e\<^sub>p e cd st = Normal ((), st6')" - then show "gas st6' \ gas st" - proof cases - assume "gas st \ g" - with 39 stmt_def show ?thesis using stmt.psimps(4) g_def by simp - next - assume "\ gas st \ g" - show ?thesis - proof (cases "expr ex e\<^sub>p e cd (st\gas := gas st - g\)") - case (n a st') - then show ?thesis - proof (cases a) - case (Pair b c) - then show ?thesis - proof (cases b) - case (KValue b) - then show ?thesis - proof (cases c) - case (Value x1) - then show ?thesis - proof (cases x1) - case (TSInt x1) - with 39(1) stmt_def `\ gas st \ g` n Pair KValue Value show ?thesis using stmt.psimps(4) g_def by simp - next - case (TUInt x2) - with 39(1) stmt_def `\ gas st \ g` n Pair KValue Value show ?thesis using stmt.psimps(4) g_def by simp - next - case TBool - then show ?thesis - proof cases - assume "b = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True" - with 39(1) `\ gas st \ g` n Pair KValue Value TBool have "stmt (ITE ex s1 s2) e\<^sub>p e cd st = stmt s1 e\<^sub>p e cd st'" using stmt.psimps(4) g_def by simp - with 39(3) stmt_def `\ gas st \ g` n Pair KValue Value TBool `b = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True` have "gas st6' \ gas st'" using g_def by simp - moreover from 39(2) `\ gas st \ g` n Pair have "gas st' \ gas st" using g_def by simp - ultimately show ?thesis by arith - next - assume "\ b = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True" - with 39(1) `\ gas st \ g` n Pair KValue Value TBool have "stmt (ITE ex s1 s2) e\<^sub>p e cd st = stmt s2 e\<^sub>p e cd st'" using stmt.psimps(4) g_def by simp - with 39(4) stmt_def `\ gas st \ g` n Pair KValue Value TBool `\ b = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True` have "gas st6' \ gas st'" using g_def by simp - moreover from 39(2) `\ gas st \ g` n Pair have "gas st' \ gas st" using g_def by simp - ultimately show ?thesis by arith - qed - next - case TAddr - with 39(1) stmt_def `\ gas st \ g` n Pair KValue Value show ?thesis using stmt.psimps(4) g_def by simp - qed - next - case (Calldata x2) - with 39(1) stmt_def `\ gas st \ g` n Pair KValue show ?thesis using stmt.psimps(4) g_def by simp - next - case (Memory x3) - with 39(1) stmt_def `\ gas st \ g` n Pair KValue show ?thesis using stmt.psimps(4) g_def by simp - next - case (Storage x4) - with 39(1) stmt_def `\ gas st \ g` n Pair KValue show ?thesis using stmt.psimps(4) g_def by simp - qed - next - case (KCDptr x2) - with 39(1) stmt_def `\ gas st \ g` n Pair show ?thesis using stmt.psimps(4) g_def by simp - next - case (KMemptr x3) - with 39(1) stmt_def `\ gas st \ g` n Pair show ?thesis using stmt.psimps(4) g_def by simp - next - case (KStoptr x4) - with 39(1) stmt_def `\ gas st \ g` n Pair show ?thesis using stmt.psimps(4) g_def by simp - qed - qed - next - case (e e) - with 39(1) stmt_def `\ gas st \ g` show ?thesis using stmt.psimps(4) g_def by simp - qed - qed - qed -next - case (40 ex s0 e\<^sub>p e cd st) - define g where "g = costs (WHILE ex s0) e\<^sub>p e cd st" - show ?case - proof (rule allI[OF impI]) - fix st6' - assume stmt_def: "stmt (WHILE ex s0) e\<^sub>p e cd st = Normal ((), st6')" - then show "gas st6' \ gas st" - proof cases - assume "gas st \ costs (WHILE ex s0) e\<^sub>p e cd st" - with 40(1) stmt_def show ?thesis using stmt.psimps(5) by simp - next - assume gcost: "\ gas st \ costs (WHILE ex s0) e\<^sub>p e cd st" - show ?thesis - proof (cases "expr ex e\<^sub>p e cd (st\gas := gas st - g\)") - case (n a st') - then show ?thesis - proof (cases a) - case (Pair b c) - then show ?thesis - proof (cases b) - case (KValue b) - then show ?thesis - proof (cases c) - case (Value x1) - then show ?thesis - proof (cases x1) - case (TSInt x1) - with 40(1) stmt_def gcost n Pair KValue Value show ?thesis using stmt.psimps(5) g_def by simp - next - case (TUInt x2) - with 40(1) stmt_def gcost n Pair KValue Value show ?thesis using stmt.psimps(5) g_def by simp - next - case TBool - then show ?thesis - proof cases - assume "b = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True" - then show ?thesis - proof (cases "stmt s0 e\<^sub>p e cd st'") - case n2: (n a st'') - with 40(1) gcost n Pair KValue Value TBool `b = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True` have "stmt (WHILE ex s0) e\<^sub>p e cd st = stmt (WHILE ex s0) e\<^sub>p e cd st''" using stmt.psimps(5)[of ex s0 e\<^sub>p e cd st] g_def by (simp add: Let_def split:unit.split_asm) - with 40(4) stmt_def gcost n2 Pair KValue Value TBool `b = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True` n have "gas st6' \ gas st''" using g_def by simp - moreover from 40(3) gcost n2 Pair KValue Value TBool `b = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True` n have "gas st'' \ gas st'" using g_def by simp - moreover from 40(2)[of _ "st\gas := gas st - g\"] gcost n Pair have "gas st' \ gas st" using g_def by simp - ultimately show ?thesis by simp - next - case (e x) - with 40(1) stmt_def gcost n Pair KValue Value TBool `b = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True` show ?thesis using stmt.psimps(5) g_def by (simp split: Ex.split_asm) - qed - next - assume "\ b = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True" - with 40(1) gcost n Pair KValue Value TBool have "stmt (WHILE ex s0) e\<^sub>p e cd st = return () st'" using stmt.psimps(5) g_def by simp - with stmt_def have "gas st6' \ gas st'" by simp - moreover from 40(2)[of _ "st\gas := gas st - g\"] gcost n have "gas st' \ gas st" using g_def by simp - ultimately show ?thesis by simp - qed - next - case TAddr - with 40(1) stmt_def gcost n Pair KValue Value show ?thesis using stmt.psimps(5) g_def by simp - qed - next - case (Calldata x2) - with 40(1) stmt_def gcost n Pair KValue show ?thesis using stmt.psimps(5) g_def by simp - next - case (Memory x3) - with 40(1) stmt_def gcost n Pair KValue show ?thesis using stmt.psimps(5) g_def by simp - next - case (Storage x4) - with 40(1) stmt_def gcost n Pair KValue show ?thesis using stmt.psimps(5) g_def by simp - qed - next - case (KCDptr x2) - with 40(1) stmt_def gcost n Pair show ?thesis using stmt.psimps(5) g_def by simp - next - case (KMemptr x3) - with 40(1) stmt_def gcost n Pair show ?thesis using stmt.psimps(5) g_def by simp - next - case (KStoptr x4) - with 40(1) stmt_def gcost n Pair show ?thesis using stmt.psimps(5) g_def by simp - qed - qed - next - case (e e) - with 40(1) stmt_def gcost show ?thesis using stmt.psimps(5) g_def by simp - qed - qed - qed -next - case (41 i xe e\<^sub>p e cd st) - define g where "g = costs (INVOKE i xe) e\<^sub>p e cd st" - show ?case - proof (rule allI[OF impI]) - fix st6' assume a1: "stmt (INVOKE i xe) e\<^sub>p e cd st = Normal ((), st6')" - show "gas st6' \ gas st" - proof (cases) - assume "gas st \ costs (INVOKE i xe) e\<^sub>p e cd st" - with 41(1) a1 show ?thesis using stmt.psimps(6) by simp - next - assume gcost: "\ gas st \ costs (INVOKE i xe) e\<^sub>p e cd st" - then show ?thesis - proof (cases "fmlookup e\<^sub>p (address e)") - case None - with 41(1) a1 gcost show ?thesis using stmt.psimps(6) by simp - next - case (Some x) - then show ?thesis - proof (cases x) - case (Pair ct _) - then show ?thesis - proof (cases "fmlookup ct i") - case None - with 41(1) g_def a1 gcost Some Pair show ?thesis using stmt.psimps(6) by simp - next - case s1: (Some a) - then show ?thesis - proof (cases a) - case (Method x1) - then show ?thesis - proof (cases x1) - case (fields fp f c) - then show ?thesis - proof (cases c) - case None - define st' e' - where "st' = st\gas := gas st - g\\stack:=emptyStore\" - and "e' = ffold (init ct) (emptyEnv (address e) (sender e) (svalue e)) (fmdom ct)" - then show ?thesis - proof (cases "load False fp xe e\<^sub>p e' emptyStore st' e cd (st\gas := gas st - g\)") - case s3: (n a st''') - then show ?thesis - proof (cases a) - case f1: (fields e'' cd' st'') - then show ?thesis - proof (cases "stmt f e\<^sub>p e'' cd' st''") - case n2: (n a st'''') - with 41(1) g_def a1 gcost Some Pair s1 Method fields None st'_def e'_def s3 f1 - have "stmt (INVOKE i xe) e\<^sub>p e cd st = Normal ((), st''''\stack:=stack st''', memory := memory st'''\)" and *: "gas st' \ gas (st\gas := gas st - g\)" - using stmt.psimps(6)[of i xe e\<^sub>p e cd st] by (auto simp add:Let_def split:unit.split_asm) - with a1 have "gas st6' \ gas st''''" by auto - also from 41(3) gcost g_def Some Pair s1 Method fields None st'_def e'_def s3 f1 n2 - have "\ \ gas st''" by (auto split:unit.split_asm) - also have "\ \ gas st'" - proof - - from g_def gcost have "(applyf (costs (INVOKE i xe) e\<^sub>p e cd) \ (\g. assert Gas (\st. gas st \ g) (modify (\st. st\gas := gas st - g\)))) st = Normal ((), st\gas := gas st - g\)" by simp - moreover from e'_def have "e' = ffold_init ct (emptyEnv (address e) (sender e) (svalue e)) (fmdom ct)" by simp - moreover from st'_def have "applyf (\st. st\stack := emptyStore\) (st\gas := gas st - g\) = Normal (st', st\gas := gas st - g\)" by simp - ultimately have "\ev cda sta st'a. load False fp xe e\<^sub>p e' emptyStore st' e cd (st\gas := gas st - g\) = Normal ((ev, cda, sta), st'a) \ gas sta \ gas st' \ gas st'a \ gas (st\gas := gas st - g\) \ address ev = address e'" using a1 g_def gcost Some Pair s1 Method fields None st'_def e'_def s3 f1 41(2)[of _ "st\gas := gas st - g\" x ct _ _ x1 fp _ f c e' st' "st\gas := gas st - g\"] by blast - then show ?thesis using s3 f1 by auto - qed - also from * have "\ \ gas (st\gas := gas st - g\)" . - finally show ?thesis by simp - next - case (e x) - with 41(1) g_def a1 gcost Some Pair s1 Method fields None st'_def e'_def s3 f1 show ?thesis using stmt.psimps(6)[of i xe e\<^sub>p e cd st] by auto - qed - qed - next - case n2: (e x) - with 41(1) g_def a1 gcost Some Pair s1 Method fields None st'_def e'_def show ?thesis using stmt.psimps(6) by auto - qed - next - case s2: (Some a) - with 41(1) g_def a1 gcost Some Pair s1 Method fields show ?thesis using stmt.psimps(6) by simp - qed - qed - next - case (Var x2) - with 41(1) g_def a1 gcost Some Pair s1 show ?thesis using stmt.psimps(6) by simp - qed - qed - qed - qed - qed - qed -next - case (42 ad i xe val e\<^sub>p e cd st) - define g where "g = costs (EXTERNAL ad i xe val) e\<^sub>p e cd st" - show ?case - proof (rule allI[OF impI]) - fix st6' assume a1: "stmt (EXTERNAL ad i xe val) e\<^sub>p e cd st = Normal ((), st6')" - show "gas st6' \ gas st" - proof (cases) - assume "gas st \ costs (EXTERNAL ad i xe val) e\<^sub>p e cd st" - with 42(1) a1 show ?thesis using stmt.psimps(7) by simp - next - assume gcost: "\ gas st \ costs (EXTERNAL ad i xe val) e\<^sub>p e cd st" - then show ?thesis - proof (cases "expr ad e\<^sub>p e cd (st\gas := gas st - g\)") - case (n a st') - then show ?thesis - proof (cases a) - case (Pair b c) - then show ?thesis - proof (cases b) - case (KValue adv) - then show ?thesis - proof (cases c) - case (Value x1) - then show ?thesis - proof (cases x1) - case (TSInt x1) - with 42(1) g_def a1 gcost n Pair KValue Value show ?thesis using stmt.psimps(7) by auto - next - case (TUInt x2) - with 42(1) g_def a1 gcost n Pair KValue Value show ?thesis using stmt.psimps(7) by simp - next - case TBool - with 42(1) g_def a1 gcost n Pair KValue Value show ?thesis using stmt.psimps(7) by simp - next - case TAddr - then show ?thesis - proof (cases "fmlookup e\<^sub>p adv") - case None - with 42(1) g_def a1 gcost n Pair KValue Value TAddr show ?thesis using stmt.psimps(7) by simp - next - case (Some x) - then show ?thesis - proof (cases x) - case p2: (Pair ct fb) - then show ?thesis - proof (cases "expr val e\<^sub>p e cd st'") - case n1: (n kv st'') - then show ?thesis - proof (cases kv) - case p3: (Pair a b) - then show ?thesis - proof (cases a) - case k2: (KValue v) - then show ?thesis - proof (cases b) - case v: (Value t) - show ?thesis - proof (cases "fmlookup ct i") - case None - show ?thesis - proof (cases "transfer (address e) adv v (accounts st'')") - case n2: None - with 42(1) g_def a1 gcost n Pair KValue Value TAddr Some p2 None n1 p3 k2 v show ?thesis using stmt.psimps(7)[of ad i xe val e\<^sub>p e cd st] by simp - next - case s4: (Some acc) - show ?thesis - proof (cases "stmt fb e\<^sub>p (emptyEnv adv (address e) v) cd (st''\accounts := acc,stack:=emptyStore, memory:=emptyStore\)") - case n2: (n a st''') - with 42(1) g_def a1 gcost n Pair KValue Value TAddr Some p2 None n1 p3 k2 v s4 - have "stmt (EXTERNAL ad i xe val) e\<^sub>p e cd st = Normal ((), st'''\stack:=stack st'', memory := memory st''\)" - using stmt.psimps(7)[of ad i xe val e\<^sub>p e cd st] by (auto simp add:Let_def split:unit.split_asm) - with a1 have "gas st6' \ gas st'''" by auto - also from 42(6)[OF _ n Pair KValue Value TAddr Some p2 n1 p3 k2 v None _ s4, of _ st'' st'' st'' "()"] n2 g_def gcost - have "\ \ gas (st''\accounts := acc,stack:=emptyStore, memory:=emptyStore\)" by auto - also from 42(3)[of _ "st\gas := gas st - g\" _ st' _ _ adv _ x ct] g_def a1 gcost n Pair KValue Value TAddr Some p2 None n1 p3 k2 v s4 n2 - have "\ \ gas st'" by auto - also from 42(2)[of _ "st\gas := gas st - g\"] g_def a1 gcost n Pair KValue Value TAddr Some p2 None n1 p3 k2 v s4 n2 - have "\ \ gas (st\gas := gas st - g\)" by auto - finally show ?thesis by simp - next - case (e x) - with 42(1) g_def a1 gcost n Pair KValue Value TAddr Some p2 None n1 p3 k2 v s4 show ?thesis using stmt.psimps(7)[of ad i xe val e\<^sub>p e cd st] by simp - qed - qed - next - case s1: (Some a) - then show ?thesis - proof (cases a) - case (Method x1) - then show ?thesis - proof (cases x1) - case (fields fp f c) - then show ?thesis - proof (cases c) - case None - define stl e' - where "stl = st''\stack:=emptyStore, memory:=emptyStore\" - and "e' = ffold (init ct) (emptyEnv adv (address e) v) (fmdom ct)" - then show ?thesis - proof (cases "load True fp xe e\<^sub>p e' emptyStore stl e cd st''") - case s3: (n a st''') - then show ?thesis - proof (cases a) - case f1: (fields e'' cd' st'''') - show ?thesis - proof (cases "transfer (address e) adv v (accounts st'''')") - case n2: None - with 42(1) g_def a1 gcost n Pair KValue Value TAddr Some p2 s1 Method fields None n1 p3 k2 v s3 f1 stl_def e'_def show ?thesis using stmt.psimps(7)[of ad i xe val e\<^sub>p e cd st] by simp - next - case s4: (Some acc) - show ?thesis - proof (cases "stmt f e\<^sub>p e'' cd' (st''''\accounts := acc\)") - case n2: (n a st''''') - with 42(1) g_def a1 gcost n Pair KValue Value TAddr Some p2 s1 Method fields None n1 p3 k2 v stl_def e'_def s3 f1 s4 - have "stmt (EXTERNAL ad i xe val) e\<^sub>p e cd st = Normal ((), st'''''\stack:=stack st''', memory := memory st'''\)" - using stmt.psimps(7)[of ad i xe val e\<^sub>p e cd st] by (auto simp add:Let_def split:unit.split_asm) - with a1 have "gas st6' \ gas (st''''')" by auto - also from 42(5)[OF _ n Pair KValue Value TAddr Some p2 n1 p3 k2 v s1 Method fields _ None _ _ s3 _ _ _ _, of "()" f e'' "(cd', st'''')" cd' st'''' st''' st''' acc "()" "st''''\accounts := acc\"] s4 stl_def e'_def f1 n2 g_def gcost - have "\ \ gas (st''''\accounts := acc\)" by auto - also have "\ \ gas stl" - proof - - from g_def gcost have "(applyf (costs (EXTERNAL ad i xe val) e\<^sub>p e cd) \ (\g. assert Gas (\st. gas st \ g) (modify (\st. st\gas := gas st - g\)))) st = Normal ((), st\gas := gas st - g\)" by simp - moreover from e'_def have "e' = ffold_init ct (emptyEnv adv (address e) v) (fmdom ct)" by simp - moreover from n1 have "expr val e\<^sub>p e cd st' = Normal (kv, st'')" by simp - moreover from stl_def have "applyf (\st. st\stack := emptyStore, memory := emptyStore\) st'' = Normal (stl, st'')" by simp - moreover have "applyf accounts st'' = Normal ((accounts st''), st'')" by simp - ultimately have "\ev cda sta st'a. load True fp xe e\<^sub>p e' emptyStore stl e cd st'' = Normal ((ev, cda, sta), st'a) \ gas sta \ gas stl \ gas st'a \ gas st'' \ address ev = address e'" using 42(4)[of _ "st\gas := gas st - g\" _ st' _ _ adv _ x ct _ _ st'' _ b v t _ x1 fp "(f,c)" f c e'] g_def a1 gcost n Pair KValue Value TAddr Some p2 s1 Method fields None n1 p3 k2 v stl_def e'_def s3 f1 s4 n2 by blast - then show ?thesis using s3 f1 by auto - qed - also from stl_def have "\ \ gas st''" by simp - also from 42(3)[of _ "st\gas := gas st - g\" _ st' _ _ adv _ x ct] g_def a1 gcost n Pair KValue Value TAddr Some p2 s1 Method fields None n1 p3 k2 v stl_def e'_def s3 f1 s4 n2 - have "\ \ gas st'" by auto - also from 42(2)[of _ "st\gas := gas st - g\"] g_def a1 gcost n Pair KValue Value TAddr Some p2 s1 Method fields None n1 p3 k2 v stl_def e'_def s3 f1 s4 n2 - have "\ \ gas (st\gas := gas st - g\)" by auto - finally show ?thesis by simp - next - case (e x) - with 42(1) g_def a1 gcost n Pair KValue Value TAddr Some p2 s1 Method fields None n1 p3 k2 v stl_def e'_def s3 f1 s4 show ?thesis using stmt.psimps(7)[of ad i xe val e\<^sub>p e cd st] by simp - qed - qed - qed - next - case (e x) - with 42(1) g_def a1 gcost n Pair KValue Value TAddr Some p2 s1 Method fields None n1 p3 k2 v stl_def e'_def show ?thesis using stmt.psimps(7)[of ad i xe val e\<^sub>p e cd st] by simp - qed - next - case s2: (Some a) - with 42(1) g_def a1 gcost n Pair KValue Value TAddr Some p2 s1 Method fields n1 p3 k2 v show ?thesis using stmt.psimps(7)[of ad i xe val e\<^sub>p e cd st] by simp - qed - qed - next - case (Var x2) - with 42(1) g_def a1 gcost n Pair KValue Value TAddr Some p2 s1 n1 p3 k2 v show ?thesis using stmt.psimps(7)[of ad i xe val e\<^sub>p e cd st] by simp - qed - qed - next - case (Calldata x2) - with 42(1) g_def a1 gcost n Pair KValue Value TAddr Some p2 n1 p3 k2 show ?thesis using stmt.psimps(7)[of ad i xe val e\<^sub>p e cd st] by simp - next - case (Memory x3) - with 42(1) g_def a1 gcost n Pair KValue Value TAddr Some p2 n1 p3 k2 show ?thesis using stmt.psimps(7)[of ad i xe val e\<^sub>p e cd st] by simp - next - case (Storage x4) - with 42(1) g_def a1 gcost n Pair KValue Value TAddr Some p2 n1 p3 k2 show ?thesis using stmt.psimps(7)[of ad i xe val e\<^sub>p e cd st] by simp - qed - next - case (KCDptr x2) - with 42(1) g_def a1 gcost n Pair KValue Value TAddr Some p2 n1 p3 show ?thesis using stmt.psimps(7)[of ad i xe val e\<^sub>p e cd st] by simp - next - case (KMemptr x3) - with 42(1) g_def a1 gcost n Pair KValue Value TAddr Some p2 n1 p3 show ?thesis using stmt.psimps(7)[of ad i xe val e\<^sub>p e cd st] by simp - next - case (KStoptr x4) - with 42(1) g_def a1 gcost n Pair KValue Value TAddr Some p2 n1 p3 show ?thesis using stmt.psimps(7)[of ad i xe val e\<^sub>p e cd st] by simp - qed - qed - next - case n2: (e x) - with 42(1) g_def a1 gcost n Pair KValue Value TAddr Some p2 show ?thesis using stmt.psimps(7)[of ad i xe val e\<^sub>p e cd st] by simp - qed - qed - qed - qed - next - case (Calldata x2) - with 42(1) g_def a1 gcost n Pair KValue show ?thesis using stmt.psimps(7)[of ad i xe val e\<^sub>p e cd st] by simp - next - case (Memory x3) - with 42(1) g_def a1 gcost n Pair KValue show ?thesis using stmt.psimps(7)[of ad i xe val e\<^sub>p e cd st] by simp - next - case (Storage x4) - with 42(1) g_def a1 gcost n Pair KValue show ?thesis using stmt.psimps(7)[of ad i xe val e\<^sub>p e cd st] by simp - qed - next - case (KCDptr x2) - with 42(1) g_def a1 gcost n Pair show ?thesis using stmt.psimps(7)[of ad i xe val e\<^sub>p e cd st] by simp - next - case (KMemptr x3) - with 42(1) g_def a1 gcost n Pair show ?thesis using stmt.psimps(7)[of ad i xe val e\<^sub>p e cd st] by simp - next - case (KStoptr x4) - with 42(1) g_def a1 gcost n Pair show ?thesis using stmt.psimps(7)[of ad i xe val e\<^sub>p e cd st] by simp - qed - qed - next - case (e _) - with 42(1) g_def a1 gcost show ?thesis using stmt.psimps(7)[of ad i xe val e\<^sub>p e cd st] by simp - qed - qed - qed -next - case (43 ad ex e\<^sub>p e cd st) - define g where "g = costs (TRANSFER ad ex) e\<^sub>p e cd st" - show ?case - proof (rule allI[OF impI]) - fix st6' assume stmt_def: "stmt (TRANSFER ad ex) e\<^sub>p e cd st = Normal ((), st6')" - show "gas st6' \ gas st" - proof cases - assume "gas st \ g" - with 43 stmt_def g_def show ?thesis using stmt.psimps(8)[of ad ex e\<^sub>p e cd st] by simp - next - assume "\ gas st \ g" - show ?thesis - proof (cases "expr ex e\<^sub>p e cd (st\gas := gas st - g\)") - case (n a st') - then show ?thesis - proof (cases a) - case (Pair b c) - then show ?thesis - proof (cases b) - case (KValue v) - then show ?thesis - proof (cases c) - case (Value t) - then show ?thesis - proof (cases "expr ad e\<^sub>p e cd st'") - case n2: (n a st'') - then show ?thesis - proof (cases a) - case p2: (Pair b c) - then show ?thesis - proof (cases b) - case k2: (KValue adv) - then show ?thesis - proof (cases c) - case v2: (Value x1) - then show ?thesis - proof (cases x1) - case (TSInt x1) - with 43(1) stmt_def `\ gas st \ g` n Pair KValue Value n2 p2 k2 v2 g_def show ?thesis using stmt.psimps(8) by simp - next - case (TUInt x2) - with 43(1) stmt_def `\ gas st \ g` n Pair KValue Value n2 p2 k2 v2 g_def show ?thesis using stmt.psimps(8) by simp - next - case TBool - with 43(1) stmt_def `\ gas st \ g` n Pair KValue Value n2 p2 k2 v2 g_def show ?thesis using stmt.psimps(8) by simp - next - case TAddr - then show ?thesis - proof (cases "transfer (address e) adv v (accounts st'')") - case None - with 43(1) stmt_def g_def `\ gas st \ g` n Pair KValue Value n2 p2 k2 v2 TAddr show ?thesis using stmt.psimps(8) by simp - next - case (Some acc) - then show ?thesis - proof (cases "fmlookup e\<^sub>p adv") - case None - with 43(1) `\ gas st \ g` n Pair KValue Value n2 p2 k2 v2 TAddr Some g_def - have "stmt (TRANSFER ad ex) e\<^sub>p e cd st = Normal ((),st''\accounts:=acc\)" using stmt.psimps(8)[of ad ex e\<^sub>p e cd st] by simp - with stmt_def have "gas st6' \ gas st''" by auto - also from 43(3)[of "()" "(st\gas := gas st - g\)" _ st'] `\ gas st \ g` n Pair KValue Value n2 g_def have "\ \ gas st'" by simp - also from 43(2)[of "()" "(st\gas := gas st - g\)"] `\ gas st \ g` n g_def have "\ \ gas st" by simp - finally show ?thesis . - next - case s2: (Some a) - then show ?thesis - proof (cases a) - case p3: (Pair ct f) - define e' where "e' = ffold_init ct (emptyEnv adv (address e) v) (fmdom ct)" - show ?thesis - proof (cases "stmt f e\<^sub>p e' emptyStore (st''\accounts := acc, stack:=emptyStore, memory:=emptyStore\)") - case n3: (n a st''') - with 43(1) `\ gas st \ g` n Pair KValue Value n2 p2 k2 v2 TAddr Some s2 p3 g_def - have "stmt (TRANSFER ad ex) e\<^sub>p e cd st = Normal ((),st'''\stack:=stack st'', memory := memory st''\)" using e'_def stmt.psimps(8)[of ad ex e\<^sub>p e cd st] by simp - with stmt_def have "gas st6' \ gas st'''" by auto - also from 43(4)[of "()" "st\gas := gas st - g\" _ st' _ _ v t _ st'' _ _ adv x1 "accounts st''" st'' acc _ ct f e' _ st'' "()" "st''\accounts := acc, stack:=emptyStore, memory:=emptyStore\"] `\ gas st \ g` n Pair KValue Value n2 p2 k2 v2 TAddr Some s2 p3 g_def n2 e'_def n3 - have "\ \ gas (st''\accounts := acc, stack := emptyStore, memory := emptyStore\)" by simp - also from 43(3)[of "()" "(st\gas := gas st - g\)" _ st'] `\ gas st \ g` n Pair KValue Value n2 g_def have "\ \ gas st'" by simp - also from 43(2)[of "()" "(st\gas := gas st - g\)"] `\ gas st \ g` n g_def have "\ \ gas st" by simp - finally show ?thesis . - next - case (e x) - with 43(1) `\ gas st \ g` n Pair KValue Value n2 p2 k2 v2 TAddr Some s2 p3 g_def e'_def stmt_def show ?thesis using stmt.psimps(8)[of ad ex e\<^sub>p e cd st] by simp - qed - qed - qed - qed - qed - next - case (Calldata x2) - with 43(1) stmt_def `\ gas st \ g` n Pair KValue Value n2 p2 k2 g_def show ?thesis using stmt.psimps(8) by simp - next - case (Memory x3) - with 43(1) stmt_def `\ gas st \ g` n Pair KValue Value n2 p2 k2 g_def show ?thesis using stmt.psimps(8) by simp - next - case (Storage x4) - with 43(1) stmt_def `\ gas st \ g` n Pair KValue Value n2 p2 k2 g_def show ?thesis using stmt.psimps(8) by simp - qed - next - case (KCDptr x2) - with 43(1) stmt_def `\ gas st \ g` n Pair KValue Value n2 p2 g_def show ?thesis using stmt.psimps(8) by simp - next - case (KMemptr x3) - with 43(1) stmt_def `\ gas st \ g` n Pair KValue Value n2 p2 g_def show ?thesis using stmt.psimps(8) by simp - next - case (KStoptr x4) - with 43(1) stmt_def `\ gas st \ g` n Pair KValue Value n2 p2 g_def show ?thesis using stmt.psimps(8) by simp - qed - qed - next - case (e e) - with 43(1) stmt_def `\ gas st \ g` n Pair KValue Value g_def show ?thesis using stmt.psimps(8) by simp - qed - next - case (Calldata x2) - with 43(1) stmt_def `\ gas st \ g` n Pair KValue g_def show ?thesis using stmt.psimps(8) by simp - next - case (Memory x3) - with 43(1) stmt_def `\ gas st \ g` n Pair KValue g_def show ?thesis using stmt.psimps(8) by simp - next - case (Storage x4) - with 43(1) stmt_def `\ gas st \ g` n Pair KValue g_def show ?thesis using stmt.psimps(8) by simp - qed - next - case (KCDptr x2) - with 43(1) stmt_def `\ gas st \ g` n Pair g_def show ?thesis using stmt.psimps(8) by simp - next - case (KMemptr x3) - with 43(1) stmt_def `\ gas st \ g` n Pair g_def show ?thesis using stmt.psimps(8) by simp - next - case (KStoptr x4) - with 43(1) stmt_def `\ gas st \ g` n Pair g_def show ?thesis using stmt.psimps(8) by simp - qed - qed - next - case (e e) - with 43(1) stmt_def `\ gas st \ g` g_def show ?thesis using stmt.psimps(8) by simp - qed - qed - qed -next - case (44 id0 tp ex s e\<^sub>p e\<^sub>v cd st) - define g where "g = costs (BLOCK ((id0, tp), ex) s) e\<^sub>p e\<^sub>v cd st" - show ?case - proof (rule allI[OF impI]) - fix st6' assume stmt_def: "stmt (BLOCK ((id0, tp), ex) s) e\<^sub>p e\<^sub>v cd st = Normal ((), st6')" - show "gas st6' \ gas st" - proof cases - assume "gas st \ g" - with 44 stmt_def g_def show ?thesis using stmt.psimps(9) by simp - next - assume "\ gas st \ g" - show ?thesis - proof (cases ex) - case None - then show ?thesis - proof (cases "decl id0 tp None False cd (memory (st\gas := gas st - g\)) cd e\<^sub>v (st\gas := gas st - g\)") - case (n a st') - then show ?thesis - proof (cases a) - case (Pair cd' e') - with 44(1) stmt_def `\ gas st \ g` None n g_def have "stmt (BLOCK ((id0, tp), ex) s) e\<^sub>p e\<^sub>v cd st = stmt s e\<^sub>p e' cd' st'" using stmt.psimps(9)[of id0 tp ex s e\<^sub>p e\<^sub>v cd st] by (simp split:unit.split_asm) - with 44(4)[of "()" "st\gas := gas st - g\"] stmt_def `\ gas st \ g` None n Pair g_def have "gas st6' \ gas st'" by simp - moreover from n Pair have "gas st' \ gas st" using decl_gas_address by simp - ultimately show ?thesis by simp - qed - next - case (e x) - with 44 stmt_def `\ gas st \ g` None g_def show ?thesis using stmt.psimps(9) by simp - qed - next - case (Some ex') - then show ?thesis - proof (cases "expr ex' e\<^sub>p e\<^sub>v cd (st\gas := gas st - g\)") - case (n a st') - then show ?thesis - proof (cases a) - case (Pair v t) - then show ?thesis - proof (cases "decl id0 tp (Some (v, t)) False cd (memory st') cd e\<^sub>v st'") - case s2: (n a st'') - then show ?thesis - proof (cases a) - case f2: (Pair cd' e') - with 44(1) stmt_def `\ gas st \ g` Some n Pair s2 g_def have "stmt (BLOCK ((id0, tp), ex) s) e\<^sub>p e\<^sub>v cd st = stmt s e\<^sub>p e' cd' st''" using stmt.psimps(9)[of id0 tp ex s e\<^sub>p e\<^sub>v cd st] by (simp split:unit.split_asm) - with 44(3)[of "()" "st\gas := gas st - g\" ex' _ st'] stmt_def `\ gas st \ g` Some n Pair s2 f2 g_def have "gas st6' \ gas st''" by simp - moreover from Some n Pair s2 f2 g_def have "gas st'' \ gas st'" using decl_gas_address by simp - moreover from 44(2)[of "()" "st\gas := gas st - g\" ex'] `\ gas st \ g` Some n Pair g_def have "gas st' \ gas st" by simp - ultimately show ?thesis by simp - qed - next - case (e x) - with 44(1) stmt_def `\ gas st \ g` Some n Pair g_def show ?thesis using stmt.psimps(9) by simp - qed - qed - next - case (e e) - with 44 stmt_def `\ gas st \ g` Some g_def show ?thesis using stmt.psimps(9) by simp - qed - qed - qed - qed -qed - -subsection \Termination\ - -lemma x1: - assumes "expr x e\<^sub>p env cd st = Normal (val, s')" - and "msel_ssel_lexp_expr_load_rexp_stmt_dom (Inr (Inl (Inl (x, e\<^sub>p, env, cd, st))))" - shows "gas s' < gas st \ gas s' = gas st" - using assms msel_ssel_lexp_expr_load_rexp_stmt_dom_gas(4)[of x e\<^sub>p env cd st] by auto - -lemma x2: -assumes "(if gas st \ c then throw Gas st else (get \ (\s. put (s\gas := gas s - c\))) st) = Normal ((), s')" - and "expr x e\<^sub>p e cd s' = Normal (val, s'a)" - and "msel_ssel_lexp_expr_load_rexp_stmt_dom (Inr (Inl (Inl (x, e\<^sub>p, e, cd, s'))))" - shows "gas s'a < gas st \ gas s'a = gas st" -proof - - from assms have "gas s' < gas st \ gas s' = gas st" by (auto split:if_split_asm) - with assms show ?thesis using x1[of x e\<^sub>p e cd s' val s'a] by auto -qed - -lemma x2sub: - assumes "(if gas st \ costs (TRANSFER ad ex) e\<^sub>p e cd st then throw Gas st - else (get \ (\s. put (s\gas := gas s - costs (TRANSFER ad ex) e\<^sub>p e cd st\))) st) = - Normal ((), s')" and -" expr ex e\<^sub>p e cd s' = Normal ((KValue vb, Value t), s'a)" -and " msel_ssel_lexp_expr_load_rexp_stmt_dom (Inr (Inl (Inl (ex, e\<^sub>p, e, cd, s'))))" -and "(\ad i xe val e\<^sub>p e cd st. 0 < costs (EXTERNAL ad i xe val) e\<^sub>p e cd st)" and "gas s'a \ gas st" shows "gas s'a < gas st" -proof - - from assms have "gas s' < gas st \ gas s' = gas st" by (auto split:if_split_asm) - with assms show ?thesis using x1[of ex e\<^sub>p e cd s' "(KValue vb, Value t)" s'a] by auto -qed - -lemma x3: - assumes "(if gas st \ c then throw Gas st else (get \ (\s. put (s\gas := gas s - c\))) st) = Normal ((), s')" - and "s'\stack := emptyStore\ = va" - and "load False ad xe e\<^sub>p(ffold (init aa) \address = address e, sender = sender e, svalue = svalue e, denvalue = fmempty\ (fmdom aa)) emptyStore va e cd s' = Normal ((ag, ah, s'd), vc)" - and "msel_ssel_lexp_expr_load_rexp_stmt_dom (Inr (Inl (Inr (False, ad, xe, e\<^sub>p, ffold (init aa) \address = address e, sender = sender e, svalue = svalue e, denvalue = fmempty\ (fmdom aa), emptyStore, va, e, cd, s'))))" - and "c>0" - shows "gas s'd < gas st" -proof - - from assms have "gas s'd \ gas va" using msel_ssel_lexp_expr_load_rexp_stmt_dom_gas(5)[of False ad xe e\<^sub>p "ffold (init aa) \address = address e, sender = sender e, svalue = svalue e, denvalue = fmempty\ (fmdom aa)" emptyStore va e cd s'] by blast - also from assms(2) have "\ = gas s'" by auto - also from assms(1,5) have "\ < gas st" by (auto split:if_split_asm) - finally show ?thesis . -qed - -lemma x4: - assumes "(if gas st \ c then throw Gas st else (get \ (\s. put (s\gas := gas s - c\))) st) = Normal ((), s')" - and "s'\stack := emptyStore\ = va" - and "load False ad xe e\<^sub>p (ffold (init aa) \address = address e, sender = sender e, svalue = svalue e, denvalue = fmempty\ (fmdom aa)) emptyStore va e cd s' = Normal ((ag, ah, s'd), vc)" - and "stmt ae e\<^sub>p ag ah s'd = Normal ((), s'e)" - and "msel_ssel_lexp_expr_load_rexp_stmt_dom (Inr (Inr (Inr (ae, e\<^sub>p, ag, ah, s'd))))" - and "msel_ssel_lexp_expr_load_rexp_stmt_dom (Inr (Inl (Inr (False, ad, xe, e\<^sub>p, ffold (init aa) \address = address e, sender = sender e, svalue = svalue e, denvalue = fmempty\ (fmdom aa), emptyStore, va, e, cd, s'))))" - and "c>0" - shows "gas s'e < gas st" -proof - - from assms have "gas s'e \ gas s'd" using msel_ssel_lexp_expr_load_rexp_stmt_dom_gas(7) by blast - with assms show ?thesis using x3[OF assms(1) assms(2) assms(3) assms(6)] by simp -qed - -lemma x5: - assumes "(if gas st \ costs (COMP s1 s2) e\<^sub>p e cd st then throw Gas st else (get \ (\s. put (s\gas := gas s - costs (COMP s1 s2) e\<^sub>p e cd st\))) st) = Normal ((), s')" - and "stmt s1 e\<^sub>p e cd s' = Normal ((), s'a)" - and "msel_ssel_lexp_expr_load_rexp_stmt_dom (Inr (Inr (Inr (s1, e\<^sub>p, e, cd, s'))))" - shows "gas s'a < gas st \ gas s'a = gas st" - using assms msel_ssel_lexp_expr_load_rexp_stmt_dom_gas(7)[of s1 e\<^sub>p e cd s'] by (auto split:if_split_asm) - -lemma x6: - assumes "(if gas st \ costs (WHILE ex s0) e\<^sub>p e cd st then throw Gas st else (get \ (\s. put (s\gas := gas s - costs (WHILE ex s0) e\<^sub>p e cd st\))) st) = Normal ((), s')" - and "expr ex e\<^sub>p e cd s' = Normal (val, s'a)" - and "stmt s0 e\<^sub>p e cd s'a = Normal ((), s'b)" - and "msel_ssel_lexp_expr_load_rexp_stmt_dom (Inr (Inr (Inr (s0, e\<^sub>p, e, cd, s'a))))" - and "msel_ssel_lexp_expr_load_rexp_stmt_dom (Inr (Inl (Inl (ex, e\<^sub>p, e, cd, s'))))" - shows "gas s'b < gas st" -proof - - from assms have "gas s'b \ gas s'a" using msel_ssel_lexp_expr_load_rexp_stmt_dom_gas(7)[of s0 e\<^sub>p e cd s'a] by blast - also from assms have "\ \ gas s'" using msel_ssel_lexp_expr_load_rexp_stmt_dom_gas(4)[of ex e\<^sub>p e cd s'] by auto - also from assms(1) have "\ < gas st" using while_not_zero by (auto split:if_split_asm) - finally show ?thesis . -qed - -lemma x7: - assumes "(if gas st \ c then throw Gas st else (get \ (\s. put (s\gas := gas s - c\))) st) = Normal ((), s')" - and "expr ad e\<^sub>p e cd s' = Normal ((KValue vb, Value TAddr), s'a)" - and "expr val e\<^sub>p e cd s'a = Normal ((KValue vd, Value ta), s'b)" - and "msel_ssel_lexp_expr_load_rexp_stmt_dom (Inr (Inl (Inl (val, e\<^sub>p, e, cd, s'a))))" - and "msel_ssel_lexp_expr_load_rexp_stmt_dom (Inr (Inl (Inl (ad, e\<^sub>p, e, cd, s'))))" - and "c>0" - shows "gas s'b < gas st" -proof - - from assms(4,3) have "gas s'b \ gas s'a" using msel_ssel_lexp_expr_load_rexp_stmt_dom_gas(4)[of val e\<^sub>p e cd s'a] by simp - also from assms(5,2) have "\ \ gas s'" using msel_ssel_lexp_expr_load_rexp_stmt_dom_gas(4)[of ad e\<^sub>p e cd s'] by simp - also from assms(1,6) have "\ < gas st" by (auto split:if_split_asm) - finally show ?thesis . -qed - -lemma x8: - assumes "(if gas st \ costs (TRANSFER ad ex) e\<^sub>p e cd st then throw Gas st else (get \ (\s. put (s\gas := gas s - costs (TRANSFER ad ex) e\<^sub>p e cd st\))) st) = Normal ((), s')" - and "expr ex e\<^sub>p e cd s' = Normal ((KValue vb, Value t), s'a)" - and "expr ad e\<^sub>p e cd s'a = Normal ((KValue vd, Value TAddr), s'b)" - and "s'b\accounts := ab, stack := emptyStore, memory := emptyStore\ = s'e" - and "msel_ssel_lexp_expr_load_rexp_stmt_dom (Inr (Inl (Inl (ad, e\<^sub>p, e, cd, s'a))))" - and "msel_ssel_lexp_expr_load_rexp_stmt_dom (Inr (Inl (Inl (ex, e\<^sub>p, e, cd, s'))))" - shows "gas s'e < gas st" -proof - - from assms(4) have "gas s'e = gas s'b" by auto - also from assms(5,3) have "\ \ gas s'a" using msel_ssel_lexp_expr_load_rexp_stmt_dom_gas(4)[of ad e\<^sub>p e cd s'a] by simp - also from assms(6,2) have "\ \ gas s'" using msel_ssel_lexp_expr_load_rexp_stmt_dom_gas(4)[of ex e\<^sub>p e cd s'] by simp - also from assms(1) have "\ < gas st" using transfer_not_zero by (auto split:if_split_asm) - finally show ?thesis . -qed - -lemma x9: - assumes "(if gas st \ costs (BLOCK ((id0, tp), Some a) s) e\<^sub>p e\<^sub>v cd st then throw Gas st else (get \ (\sa. put (sa\gas := gas sa - costs (BLOCK ((id0, tp), Some a) s) e\<^sub>p e\<^sub>v cd st\))) st) = Normal ((), s')" - and "expr a e\<^sub>p e\<^sub>v cd s' = Normal ((aa, b), s'a)" - and "decl id0 tp (Some (aa, b)) False cd vb cd e\<^sub>v s'a = Normal ((ab, ba), s'c)" - and "msel_ssel_lexp_expr_load_rexp_stmt_dom (Inr (Inl (Inl (a, e\<^sub>p, e\<^sub>v, cd, s'))))" - shows "gas s'c < gas st \ gas s'c = gas st" -proof - - from assms have "gas s'c = gas s'a" using decl_gas_address[of id0 tp "(Some (aa, b))" False cd vb cd e\<^sub>v s'a] by simp - also from assms have "\ \ gas s'" using msel_ssel_lexp_expr_load_rexp_stmt_dom_gas(4)[of a e\<^sub>p e\<^sub>v cd s'] by simp - also from assms(1) have "\ \ gas st" by (auto split:if_split_asm) - finally show ?thesis by auto -qed - -lemma x10: - assumes "(if gas st \ costs (BLOCK ((id0, tp), None) s) e\<^sub>p e\<^sub>v cd st then throw Gas st else (get \ (\sa. put (sa\gas := gas sa - costs (BLOCK ((id0, tp), None) s) e\<^sub>p e\<^sub>v cd st\))) st) = Normal ((), s')" - and "decl id0 tp None False cd va cd e\<^sub>v s' = Normal ((a, b), s'b)" - shows "gas s'b < gas st \ gas s'b = gas st" -proof - - from assms have "gas s'b = gas s'" using decl_gas_address[of id0 tp None False cd va cd e\<^sub>v s'] by simp - also from assms(1) have "\ \ gas st" by (auto split:if_split_asm) - finally show ?thesis by auto -qed - -lemma x11: - assumes "(if gas st \ c then throw Gas st else (get \ (\s. put (s\gas := gas s - c\))) st) = Normal ((), s')" - and "expr ad e\<^sub>p e cd s' = Normal ((KValue vb, Value TAddr), s'a)" - and "expr val e\<^sub>p e cd s'a = Normal ((KValue vd, Value ta), s'b)" - and "load True af xe e\<^sub>p (ffold (init ab) \address = vb, sender = address e, svalue = vd, denvalue = fmempty\ (fmdom ab)) emptyStore (s'b\stack := emptyStore, memory := emptyStore\) e cd s'b = Normal ((ak, al, s'g), vh)" - and "msel_ssel_lexp_expr_load_rexp_stmt_dom (Inr (Inl (Inr (True, af, xe, e\<^sub>p, ffold (init ab) \address = vb, sender = address e, svalue = vd, denvalue = fmempty\ (fmdom ab), emptyStore, s'b\stack := emptyStore, memory := emptyStore\, e, cd, s'b))))" - and "msel_ssel_lexp_expr_load_rexp_stmt_dom (Inr (Inl (Inl (val, e\<^sub>p, e, cd, s'a))))" - and "msel_ssel_lexp_expr_load_rexp_stmt_dom (Inr (Inl (Inl (ad, e\<^sub>p, e, cd, s'))))" - and "c>0" - shows "gas s'g < gas st" -proof - - from assms have "gas s'g \ gas (s'b\stack := emptyStore, memory := emptyStore\)" using msel_ssel_lexp_expr_load_rexp_stmt_dom_gas(5)[of True af xe e\<^sub>p "ffold (init ab) \address = vb, sender = address e, svalue = vd, denvalue = fmempty\ (fmdom ab)" emptyStore "s'b\stack := emptyStore, memory := emptyStore\" e cd s'b] by blast - also have "\ = gas s'b" by simp - also from assms have "\ < gas st" using x7[OF assms(1) assms(2) assms(3) assms(6)] by auto - finally show ?thesis . -qed - -lemma x12: - assumes "(if gas st \ c then throw Gas st else (get \ (\s. put (s\gas := gas s - c\))) st) = Normal ((), s')" - and "expr ad e\<^sub>p e cd s' = Normal ((KValue vb, Value TAddr), s'a)" - and "expr val e\<^sub>p e cd s'a = Normal ((KValue vd, Value ta), s'b)" - and "load True af xe e\<^sub>p (ffold (init ab) \address = vb, sender = address e, svalue = vd, denvalue = fmempty\ (fmdom ab)) emptyStore (s'b\stack := emptyStore, memory := emptyStore\) e cd s'b = Normal ((ak, al, s'g), vh)" - and "stmt ag e\<^sub>p ak al (s'g\accounts := ala\) = Normal ((), s'h)" - and "msel_ssel_lexp_expr_load_rexp_stmt_dom (Inr (Inl (Inr (True, af, xe, e\<^sub>p, ffold (init ab) \address = vb, sender = address e, svalue = vd, denvalue = fmempty\ (fmdom ab), emptyStore, s'b\stack := emptyStore, memory := emptyStore\, e, cd, s'b))))" - and "msel_ssel_lexp_expr_load_rexp_stmt_dom (Inr (Inl (Inl (val, e\<^sub>p, e, cd, s'a))))" - and "msel_ssel_lexp_expr_load_rexp_stmt_dom (Inr (Inl (Inl (ad, e\<^sub>p, e, cd, s'))))" - and "msel_ssel_lexp_expr_load_rexp_stmt_dom (Inr (Inr (Inr (ag, e\<^sub>p, ak, al, (s'g\accounts := ala\)))))" - and "c>0" - shows "gas s'h < gas st" -proof - - from assms have "gas s'h \ gas (s'g\accounts := ala\)" using msel_ssel_lexp_expr_load_rexp_stmt_dom_gas(7) by blast - also from assms have "\ < gas st" using x11[OF assms(1) assms(2) assms(3) assms(4)] by auto - finally show ?thesis . -qed - -termination - apply (relation - "measures [\x. case x of Inr (Inr (Inr l)) \ gas (snd (snd (snd (snd l)))) - | Inr (Inr (Inl l)) \ gas (snd (snd (snd (snd l)))) - | Inr (Inl (Inr l)) \ gas (snd (snd (snd (snd (snd (snd (snd (snd (snd l))))))))) - | Inr (Inl (Inl l)) \ gas (snd (snd (snd (snd l)))) - | Inl (Inr (Inr l)) \ gas (snd (snd (snd (snd l)))) - | Inl (Inr (Inl l)) \ gas (snd (snd (snd (snd (snd (snd l)))))) - | Inl (Inl l) \ gas (snd (snd (snd (snd (snd (snd (snd l))))))), - \x. case x of Inr (Inr (Inr l)) \ 1 - | Inr (Inr (Inl l)) \ 0 - | Inr (Inl (Inr l)) \ 0 - | Inr (Inl (Inl l)) \ 0 - | Inl (Inr (Inr l)) \ 0 - | Inl (Inr (Inl l)) \ 0 - | Inl (Inl l) \ 0, - \x. case x of Inr (Inr (Inr l)) \ size (fst l) - | Inr (Inr (Inl l)) \ size (fst l) - | Inr (Inl (Inr l)) \ size_list size (fst (snd (snd l))) - | Inr (Inl (Inl l)) \ size (fst l) - | Inl (Inr (Inr l)) \ size (fst l) - | Inl (Inr (Inl l)) \ size_list size (fst (snd (snd l))) - | Inl (Inl l) \ size_list size (fst (snd (snd (snd l))))] - ") - apply simp_all - apply (simp only: x1) - apply (simp only: x1) - apply (simp only: x1) - apply (auto split: if_split_asm)[1] - apply (auto split: if_split_asm)[1] - apply (simp only: x2) - apply (simp only: x2) - apply (auto split: if_split_asm)[1] - apply (simp only: x2) - apply (auto split: if_split_asm)[1] - apply (simp only: x2) - apply (auto split: if_split_asm)[1] - apply (simp only: x2) - apply (auto split: if_split_asm)[1] - apply (simp only: x2) - apply (auto split: if_split_asm)[1] - apply (simp only: x2) - apply (auto split: if_split_asm)[1] - apply (simp only: x2) - apply (auto split: if_split_asm)[1] - apply (auto split: if_split_asm)[1] - using call_not_zero apply (simp only: x3) - using call_not_zero apply (simp add: x4) - apply (auto split: if_split_asm)[1] - apply (simp only: x2) - using ecall_not_zero apply (simp add: x7) - using ecall_not_zero apply (auto simp add: x11)[1] - using ecall_not_zero apply (auto simp add: x12)[1] - apply (simp only: x1) - apply (auto split: if_split_asm)[1] - apply (simp only: x2) - apply (simp only: x2) - apply (simp only: x2) - apply (simp only: x2) - apply (simp only: x2) - apply (auto split: if_split_asm)[1] - apply (simp only: x5) - apply (auto split: if_split_asm)[1] - apply (simp only: x2) - apply (simp only: x2) - apply (auto split: if_split_asm)[1] - apply (simp only: x2) - apply (simp only: x6) - apply (auto split: if_split_asm)[1] - using invoke_not_zero apply (simp only: x3) - apply (auto split: if_split_asm)[1] - apply (simp only: x2) - using external_not_zero apply (simp add: x7) - using external_not_zero apply (auto simp add: x11)[1] - using external_not_zero apply (auto simp add: x7)[1] - apply (auto split: if_split_asm)[1] - apply (simp add: x2) - apply (simp add: x8) - apply (auto split: if_split_asm)[1] - apply (simp only: x9) - apply (simp only: x10) - done -end - -subsection \A minimal cost model\ - -fun costs_min :: "S\ Environment\<^sub>P \ Environment \ CalldataT \ State \ Gas" - where - "costs_min SKIP e\<^sub>p e cd st = 0" -| "costs_min (ASSIGN lv ex) e\<^sub>p e cd st = 0" -| "costs_min (COMP s1 s2) e\<^sub>p e cd st = 0" -| "costs_min (ITE ex s1 s2) e\<^sub>p e cd st = 0" -| "costs_min (WHILE ex s0) e\<^sub>p e cd st = 1" -| "costs_min (TRANSFER ad ex) e\<^sub>p e cd st = 1" -| "costs_min (BLOCK ((id0, tp), ex) s) e\<^sub>p e cd st =0" -| "costs_min (INVOKE _ _) e\<^sub>p e cd st = 1" -| "costs_min (EXTERNAL _ _ _ _) e\<^sub>p e cd st = 1" - -fun costs_ex :: "E \ Environment\<^sub>P \ Environment \ CalldataT \ State \ Gas" - where - "costs_ex (E.INT _ _) e\<^sub>p e cd st = 0" -| "costs_ex (UINT _ _) e\<^sub>p e cd st = 0" -| "costs_ex (ADDRESS _) e\<^sub>p e cd st = 0" -| "costs_ex (BALANCE _) e\<^sub>p e cd st = 0" -| "costs_ex THIS e\<^sub>p e cd st = 0" -| "costs_ex SENDER e\<^sub>p e cd st = 0" -| "costs_ex VALUE e\<^sub>p e cd st = 0" -| "costs_ex (TRUE) e\<^sub>p e cd st = 0" -| "costs_ex (FALSE) e\<^sub>p e cd st = 0" -| "costs_ex (LVAL _) e\<^sub>p e cd st = 0" -| "costs_ex (PLUS _ _) e\<^sub>p e cd st = 0" -| "costs_ex (MINUS _ _) e\<^sub>p e cd st = 0" -| "costs_ex (EQUAL _ _) e\<^sub>p e cd st = 0" -| "costs_ex (LESS _ _) e\<^sub>p e cd st = 0" -| "costs_ex (AND _ _) e\<^sub>p e cd st = 0" -| "costs_ex (OR _ _) e\<^sub>p e cd st = 0" -| "costs_ex (NOT _) e\<^sub>p e cd st = 0" -| "costs_ex (CALL _ _) e\<^sub>p e cd st = 1" -| "costs_ex (ECALL _ _ _ _) e\<^sub>p e cd st = 1" - -global_interpretation solidity: statement_with_gas costs_min costs_ex - defines stmt = "solidity.stmt" - and lexp = solidity.lexp - and expr = solidity.expr - and ssel = solidity.ssel - and rexp = solidity.rexp - and msel = solidity.msel - and load = solidity.load - by unfold_locales auto - -end +section \Statements\ +theory Statements + imports Expressions StateMonad +begin + +locale statement_with_gas = expressions_with_gas + + fixes costs :: "S \ Environment \ CalldataT \ State \ Gas" + assumes while_not_zero[termination_simp]: "\e cd st ex s0. 0 < (costs (WHILE ex s0) e cd st) " + and invoke_not_zero[termination_simp]: "\e cd st i xe. 0 < (costs (INVOKE i xe) e cd st)" + and external_not_zero[termination_simp]: "\e cd st ad i xe val. 0 < (costs (EXTERNAL ad i xe val) e cd st)" + and transfer_not_zero[termination_simp]: "\e cd st ex ad. 0 < (costs (TRANSFER ad ex) e cd st)" + and new_not_zero[termination_simp]: "\e cd st i xe val. 0 < (costs (NEW i xe val) e cd st)" +begin + +subsection \Semantics of left expressions\ + +text \We first introduce lexp.\ + +fun lexp :: "L \ Environment \ CalldataT \ State \ (LType * Type, Ex, Gas) state_monad" + where "lexp (Id i) e _ st g = + (case (denvalue e) $$ i of + Some (tp, (Stackloc l)) \ return (LStackloc l, tp) + | Some (tp, (Storeloc l)) \ return (LStoreloc l, tp) + | _ \ throw Err) g" +| "lexp (Ref i r) e cd st g = + (case (denvalue e) $$ i of + Some (tp, Stackloc l) \ + (case accessStore l (stack st) of + Some (KCDptr _) \ throw Err + | Some (KMemptr l') \ + do { + t \ (case tp of Memory t \ return t | _ \ throw Err); + (l'', t') \ msel True t l' r e cd st; + return (LMemloc l'', Memory t') + } + | Some (KStoptr l') \ + do { + t \ (case tp of Storage t \ return t | _ \ throw Err); + (l'', t') \ ssel t l' r e cd st; + return (LStoreloc l'', Storage t') + } + | Some (KValue _) \ throw Err + | None \ throw Err) + | Some (tp, Storeloc l) \ + do { + t \ (case tp of Storage t \ return t | _ \ throw Err); + (l', t') \ ssel t l r e cd st; + return (LStoreloc l', Storage t') + } + | None \ throw Err) g" + +lemma lexp_gas[rule_format]: + "\l5' t5' g5'. lexp l5 ev5 cd5 st5 g5 = Normal ((l5', t5'), g5') \ g5' \ g5" +proof (induct rule: lexp.induct[where ?P="\l5 ev5 cd5 st5 g5. (\l5' t5' g5'. lexp l5 ev5 cd5 st5 g5 = Normal ((l5', t5'), g5') \ g5' \ g5)"]) + case (1 i e uv st g) + then show ?case using lexp.simps(1) by (simp split: option.split Denvalue.split prod.split) +next + case (2 i r e cd st g) + show ?case + proof (rule allI[THEN allI, THEN allI, OF impI]) + fix st5' xa xaa + assume a1: "lexp (Ref i r) e cd st g = Normal ((st5', xa), xaa)" + then show "xaa \ g" + proof (cases "fmlookup (denvalue e) i") + case None + with a1 show ?thesis using lexp.simps(2) by simp + next + case (Some a) + then show ?thesis + proof (cases a) + case (Pair tp b) + then show ?thesis + proof (cases b) + case (Stackloc l) + then show ?thesis + proof (cases "accessStore l (stack st)") + case None + with a1 Some Pair Stackloc show ?thesis using lexp.psimps(2) by simp + next + case s2: (Some a) + then show ?thesis + proof (cases a) + case (KValue x1) + with a1 Some Pair Stackloc s2 show ?thesis using lexp.psimps(2) by simp + next + case (KCDptr x2) + with a1 Some Pair Stackloc s2 show ?thesis using lexp.psimps(2) by simp + next + case (KMemptr l') + then show ?thesis + proof (cases tp) + case (Value _) + with a1 Some Pair Stackloc s2 KMemptr show ?thesis using lexp.simps(2) by simp + next + case (Calldata _) + with a1 Some Pair Stackloc s2 KMemptr show ?thesis using lexp.simps(2) by simp + next + case (Memory t) + then show ?thesis + proof (cases "msel True t l' r e cd st g") + case (n _ _) + with 2 a1 Some Pair Stackloc s2 KMemptr Memory show ?thesis using msel_ssel_expr_load_rexp_gas(1) by (simp split: prod.split_asm) + next + case (e _) + with a1 Some Pair Stackloc s2 KMemptr Memory show ?thesis using lexp.psimps(2) by simp + qed + next + case (Storage _) + with a1 Some Pair Stackloc s2 KMemptr show ?thesis using lexp.psimps(2) by simp + qed + next + case (KStoptr l') + then show ?thesis + proof (cases tp) + case (Value _) + with a1 Some Pair Stackloc s2 KStoptr show ?thesis using lexp.psimps(2) by simp + next + case (Calldata _) + with a1 Some Pair Stackloc s2 KStoptr show ?thesis using lexp.psimps(2) by simp + next + case (Memory _) + with a1 Some Pair Stackloc s2 KStoptr show ?thesis using lexp.psimps(2) by simp + next + case (Storage t) + then show ?thesis + proof (cases "ssel t l' r e cd st g") + case (n _ _) + with a1 Some Pair Stackloc s2 KStoptr Storage show ?thesis using msel_ssel_expr_load_rexp_gas(2) by (auto split: prod.split_asm) + next + case (e _) + with a1 Some Pair Stackloc s2 KStoptr Storage show ?thesis using lexp.psimps(2) by simp + qed + qed + qed + qed + next + case (Storeloc l) + then show ?thesis + proof (cases tp) + case (Value _) + with a1 Some Pair Storeloc show ?thesis using lexp.psimps(2) by simp + next + case (Calldata _) + with a1 Some Pair Storeloc show ?thesis using lexp.psimps(2) by simp + next + case (Memory _) + with a1 Some Pair Storeloc show ?thesis using lexp.psimps(2) by simp + next + case (Storage t) + then show ?thesis + proof (cases "ssel t l r e cd st g") + case (n _ _) + with a1 Some Pair Storeloc Storage show ?thesis using msel_ssel_expr_load_rexp_gas(2) by (auto split: prod.split_asm) + next + case (e _) + with a1 Some Pair Storeloc Storage show ?thesis using lexp.psimps(2) by simp + qed + qed + qed + qed + qed + qed +qed + +subsection \Semantics of statements\ + +text \The following is a helper function to connect the gas monad with the state monad.\ + +fun + toState :: "(State \ ('a, 'e, Gas) state_monad) \ ('a, 'e, State) state_monad" where + "toState gm = (\s. case gm s (gas s) of + Normal (a,g) \ Normal(a,s\gas:=g\) + | Exception e \ Exception e)" + +lemma wptoState[wprule]: + assumes "\a g. gm s (gas s) = Normal (a, g) \ P a (s\gas:=g\)" + and "\e. gm s (gas s) = Exception e \ E e" + shows "wp (toState gm) P E s" + using assms unfolding wp_def by (simp split:result.split result.split_asm) + +text \Now we define the semantics of statements.\ +function (domintros) stmt :: "S \ Environment \ CalldataT \ (unit, Ex, State) state_monad" + where "stmt SKIP e cd st = + (do { + assert Gas (\st. gas st > costs SKIP e cd st); + modify (\st. st\gas := gas st - costs SKIP e cd st\) + }) st" +| "stmt (ASSIGN lv ex) env cd st = + (do { + assert Gas (\st. gas st > costs (ASSIGN lv ex) env cd st); + modify (\st. st\gas := gas st - costs (ASSIGN lv ex) env cd st\); + re \ toState (expr ex env cd); + case re of + (KValue v, Value t) \ + do { + rl \ toState (lexp lv env cd); + case rl of + (LStackloc l, Value t') \ + do { + v' \ option Err (\_. convert t t' v); + modify (\st. st \stack := updateStore l (KValue v') (stack st)\) + } + | (LStoreloc l, Storage (STValue t')) \ + do { + v' \ option Err (\_. convert t t' v); + modify (\st. st\storage := (storage st) (address env := fmupd l v' (storage st (address env)))\) + } + | (LMemloc l, Memory (MTValue t')) \ + do { + v' \ option Err (\_. convert t t' v); + modify (\st. st\memory := updateStore l (MValue v') (memory st)\) + } + | _ \ throw Err + } + | (KCDptr p, Calldata (MTArray x t)) \ + do { + rl \ toState (lexp lv env cd); + case rl of + (LStackloc l, Memory _) \ + do { + sv \ applyf (\st. accessStore l (stack st)); + p' \ case sv of Some (KMemptr p') \ return p' | _ \ throw Err; + m \ option Err (\st. cpm2m p p' x t cd (memory st)); + modify (\st. st\memory := m\) + } + | (LStackloc l, Storage _) \ + do { + sv \ applyf (\st. accessStore l (stack st)); + p' \ case sv of Some (KStoptr p') \ return p' | _ \ throw Err; + s \ option Err (\st. cpm2s p p' x t cd (storage st (address env))); + modify (\st. st \storage := (storage st) (address env := s)\) + } + | (LStoreloc l, _) \ + do { + s \ option Err (\st. cpm2s p l x t cd (storage st (address env))); + modify (\st. st \storage := (storage st) (address env := s)\) + } + | (LMemloc l, _) \ + do { + m \ option Err (\st. cpm2m p l x t cd (memory st)); + modify (\st. st \memory := m\) + } + | _ \ throw Err + } + | (KMemptr p, Memory (MTArray x t)) \ + do { + rl \ toState (lexp lv env cd); + case rl of + (LStackloc l, Memory _) \ modify (\st. st\stack := updateStore l (KMemptr p) (stack st)\) + | (LStackloc l, Storage _) \ + do { + sv \ applyf (\st. accessStore l (stack st)); + p' \ case sv of Some (KStoptr p') \ return p' | _ \ throw Err; + s \ option Err (\st. cpm2s p p' x t (memory st) (storage st (address env))); + modify (\st. st \storage := (storage st) (address env := s)\) + } + | (LStoreloc l, _) \ + do { + s \ option Err (\st. cpm2s p l x t (memory st) (storage st (address env))); + modify (\st. st \storage := (storage st) (address env := s)\) + } + | (LMemloc l, _) \ modify (\st. st \memory := updateStore l (MPointer p) (memory st)\) + | _ \ throw Err + } + | (KStoptr p, Storage (STArray x t)) \ + do { + rl \ toState (lexp lv env cd); + case rl of + (LStackloc l, Memory _) \ + do { + sv \ applyf (\st. accessStore l (stack st)); + p' \ case sv of Some (KMemptr p') \ return p' | _ \ throw Err; + m \ option Err (\st. cps2m p p' x t (storage st (address env)) (memory st)); + modify (\st. st\memory := m\) + } + | (LStackloc l, Storage _) \ modify (\st. st\stack := updateStore l (KStoptr p) (stack st)\) + | (LStoreloc l, _) \ + do { + s \ option Err (\st. copy p l x t (storage st (address env))); + modify (\st. st \storage := (storage st) (address env := s)\) + } + | (LMemloc l, _) \ + do { + m \ option Err (\st. cps2m p l x t (storage st (address env)) (memory st)); + modify (\st. st\memory := m\) + } + | _ \ throw Err + } + | (KStoptr p, Storage (STMap t t')) \ + do { + rl \ toState (lexp lv env cd); + l \ case rl of (LStackloc l, _) \ return l | _ \ throw Err; + modify (\st. st\stack := updateStore l (KStoptr p) (stack st)\) + } + | _ \ throw Err + }) st" +| "stmt (COMP s1 s2) e cd st = + (do { + assert Gas (\st. gas st > costs (COMP s1 s2) e cd st); + modify (\st. st\gas := gas st - costs (COMP s1 s2) e cd st\); + stmt s1 e cd; + stmt s2 e cd + }) st" +| "stmt (ITE ex s1 s2) e cd st = + (do { + assert Gas (\st. gas st > costs (ITE ex s1 s2) e cd st); + modify (\st. st\gas := gas st - costs (ITE ex s1 s2) e cd st\); + v \ toState (expr ex e cd); + b \ (case v of (KValue b, Value TBool) \ return b | _ \ throw Err); + if b = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True then stmt s1 e cd + else if b = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l False then stmt s2 e cd + else throw Err + }) st" +| "stmt (WHILE ex s0) e cd st = + (do { + assert Gas (\st. gas st > costs (WHILE ex s0) e cd st); + modify (\st. st\gas := gas st - costs (WHILE ex s0) e cd st\); + v \ toState (expr ex e cd); + b \ (case v of (KValue b, Value TBool) \ return b | _ \ throw Err); + if b = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True then + do { + stmt s0 e cd; + stmt (WHILE ex s0) e cd + } + else if b = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l False then return () + else throw Err + }) st" +| "stmt (INVOKE i xe) e cd st = + (do { + assert Gas (\st. gas st > costs (INVOKE i xe) e cd st); + modify (\st. st\gas := gas st - costs (INVOKE i xe) e cd st\); + (ct, _) \ option Err (\_. ep $$ contract e); + (fp, f) \ case ct $$ i of Some (Method (fp, False, f)) \ return (fp, f) | _ \ throw Err; + let e' = ffold_init ct (emptyEnv (address e) (contract e) (sender e) (svalue e)) (fmdom ct); + m\<^sub>o \ applyf memory; + (e\<^sub>l, cd\<^sub>l, k\<^sub>l, m\<^sub>l) \ toState (load False fp xe e' emptyStore emptyStore m\<^sub>o e cd); + k\<^sub>o \ applyf stack; + modify (\st. st\stack:=k\<^sub>l, memory:=m\<^sub>l\); + stmt f e\<^sub>l cd\<^sub>l; + modify (\st. st\stack:=k\<^sub>o\) + }) st" +(*External Method calls allow to send some money val with it*) +(*However this transfer does NOT trigger a fallback*) +(*External methods can only be called from externally*) +| "stmt (EXTERNAL ad i xe val) e cd st = + (do { + assert Gas (\st. gas st > costs (EXTERNAL ad i xe val) e cd st); + modify (\st. st\gas := gas st - costs (EXTERNAL ad i xe val) e cd st\); + kad \ toState (expr ad e cd); + adv \ case kad of (KValue adv, Value TAddr) \ return adv | _ \ throw Err; + assert Err (\_. adv \ address e); + c \ (\st. case type (accounts st adv) of Some (Contract c) \ return c st | _ \ throw Err st); + (ct, _, fb) \ option Err (\_. ep $$ c); + kv \ toState (expr val e cd); + (v, t) \ case kv of (KValue v, Value t) \ return (v, t) | _ \ throw Err; + v' \ option Err (\_. convert t (TUInt 256) v); + let e' = ffold_init ct (emptyEnv adv c (address e) v') (fmdom ct); + case ct $$ i of + Some (Method (fp, True, f)) \ + do { + (e\<^sub>l, cd\<^sub>l, k\<^sub>l, m\<^sub>l) \ toState (load True fp xe e' emptyStore emptyStore emptyStore e cd); + acc \ option Err (\st. transfer (address e) adv v' (accounts st)); + (k\<^sub>o, m\<^sub>o) \ applyf (\st. (stack st, memory st)); + modify (\st. st\accounts := acc, stack:=k\<^sub>l,memory:=m\<^sub>l\); + stmt f e\<^sub>l cd\<^sub>l; + modify (\st. st\stack:=k\<^sub>o, memory := m\<^sub>o\) + } + | None \ + do { + acc \ option Err (\st. transfer (address e) adv v' (accounts st)); + (k\<^sub>o, m\<^sub>o) \ applyf (\st. (stack st, memory st)); + modify (\st. st\accounts := acc,stack:=emptyStore, memory:=emptyStore\); + stmt fb e' emptyStore; + modify (\st. st\stack:=k\<^sub>o, memory := m\<^sub>o\) + } + | _ \ throw Err + }) st" +| "stmt (TRANSFER ad ex) e cd st = + (do { + assert Gas (\st. gas st > costs (TRANSFER ad ex) e cd st); + modify (\st. st\gas := gas st - costs (TRANSFER ad ex) e cd st\); + kv \ toState (expr ad e cd); + adv \ case kv of (KValue adv, Value TAddr) \ return adv | _ \ throw Err; + kv' \ toState (expr ex e cd); + (v, t) \ case kv' of (KValue v, Value t) \ return (v, t) | _ \ throw Err; + v' \ option Err (\_. convert t (TUInt 256) v); + acc \ applyf accounts; + case type (acc adv) of + Some (Contract c) \ + do { + (ct, _, f) \ option Err (\_. ep $$ c); + let e' = ffold_init ct (emptyEnv adv c (address e) v') (fmdom ct); + (k\<^sub>o, m\<^sub>o) \ applyf (\st. (stack st, memory st)); + acc' \ option Err (\st. transfer (address e) adv v' (accounts st)); + modify (\st. st\accounts := acc', stack:=emptyStore, memory:=emptyStore\); + stmt f e' emptyStore; + modify (\st. st\stack:=k\<^sub>o, memory := m\<^sub>o\) + } + | Some EOA \ + do { + acc' \ option Err (\st. transfer (address e) adv v' (accounts st)); + modify (\st. (st\accounts := acc'\)) + } + | None \ throw Err + }) st" +| "stmt (BLOCK ((id0, tp), None) s) e\<^sub>v cd st = + (do { + assert Gas (\st. gas st > costs (BLOCK ((id0, tp), None) s) e\<^sub>v cd st); + modify (\st. st\gas := gas st - costs (BLOCK ((id0, tp), None) s) e\<^sub>v cd st\); + (cd', mem', sck', e') \ option Err (\st. decl id0 tp None False cd (memory st) (storage st) (cd, memory st, stack st, e\<^sub>v)); + modify (\st. st\stack := sck', memory := mem'\); + stmt s e' cd' + }) st" +| "stmt (BLOCK ((id0, tp), Some ex') s) e\<^sub>v cd st = + (do { + assert Gas (\st. gas st > costs(BLOCK ((id0, tp), Some ex') s) e\<^sub>v cd st); + modify (\st. st\gas := gas st - costs (BLOCK ((id0, tp), Some ex') s) e\<^sub>v cd st\); + (v, t) \ toState (expr ex' e\<^sub>v cd); + (cd', mem', sck', e') \ option Err (\st. decl id0 tp (Some (v, t)) False cd (memory st) (storage st) (cd, memory st, stack st, e\<^sub>v)); + modify (\st. st\stack := sck', memory := mem'\); + stmt s e' cd' + }) st" +(* + Note: We cannot use (ct, (fp, cn), fb) <- option Err (\_. ep $$ i) +*) +| "stmt (NEW i xe val) e cd st = + (do { + assert Gas (\st. gas st > costs (NEW i xe val) e cd st); + modify (\st. st\gas := gas st - costs (NEW i xe val) e cd st\); + adv \ applyf (\st. hash (address e) (ShowL\<^sub>n\<^sub>a\<^sub>t (contracts (accounts st (address e))))); + assert Err (\st. type (accounts st adv) = None); + kv \ toState (expr val e cd); + (v, t) \ case kv of (KValue v, Value t) \ return (v, t) | _ \ throw Err; + (ct, cn, _) \ option Err (\_. ep $$ i); + let e' = ffold_init ct (emptyEnv adv i (address e) v) (fmdom ct); + (e\<^sub>l, cd\<^sub>l, k\<^sub>l, m\<^sub>l) \ toState (load True (fst cn) xe e' emptyStore emptyStore emptyStore e cd); + modify (\st. st\accounts := (accounts st)(adv := \bal = ShowL\<^sub>i\<^sub>n\<^sub>t 0, type = Some (Contract i), contracts = 0\), storage:=(storage st)(adv := {$$})\); + acc \ option Err (\st. transfer (address e) adv v (accounts st)); + (k\<^sub>o, m\<^sub>o) \ applyf (\st. (stack st, memory st)); + modify (\st. st\accounts := acc, stack:=k\<^sub>l, memory:=m\<^sub>l\); + stmt (snd cn) e\<^sub>l cd\<^sub>l; + modify (\st. st\stack:=k\<^sub>o, memory := m\<^sub>o\); + modify (incrementAccountContracts (address e)) + }) st" +by pat_completeness auto + +subsection \Termination\ + +text \Again, to prove termination we need a lemma regarding gas consumption.\ + +lemma stmt_dom_gas[rule_format]: + "stmt_dom (s6, ev6, cd6, st6) \ (\st6'. stmt s6 ev6 cd6 st6 = Normal((), st6') \ gas st6' \ gas st6)" +proof (induct rule: stmt.pinduct[where ?P="\s6 ev6 cd6 st6. (\st6'. stmt s6 ev6 cd6 st6 = Normal ((), st6') \ gas st6' \ gas st6)"]) + case (1 e cd st) + then show ?case using stmt.psimps(1) by simp +next + case (2 lv ex env cd st) + define g where "g = costs (ASSIGN lv ex) env cd st" + show ?case + proof (rule allI[OF impI]) + fix st6' + assume stmt_def: "stmt (ASSIGN lv ex) env cd st = Normal ((), st6')" + then show "gas st6' \ gas st" + proof cases + assume "gas st \ g" + with 2(1) stmt_def show ?thesis using stmt.psimps(2) g_def by simp + next + assume "\ gas st \ g" + define st' where "st' = st\gas := gas st - g\" + show ?thesis + proof (cases "expr ex env cd st' (gas st - g)") + case (n a g') + define st'' where "st'' = st'\gas := g'\" + then show ?thesis + proof (cases a) + case (Pair b c) + then show ?thesis + proof (cases b) + case (KValue v) + then show ?thesis + proof (cases c) + case (Value t) + then show ?thesis + proof (cases "lexp lv env cd st'' g'") + case n2: (n a g'') + then show ?thesis + proof (cases a) + case p1: (Pair a b) + then show ?thesis + proof (cases a) + case (LStackloc l) + then show ?thesis + proof (cases b) + case v2: (Value t') + then show ?thesis + proof (cases "convert t t' v ") + case None + with stmt_def `\ gas st \ g` n Pair KValue Value n2 p1 LStackloc v2 show ?thesis using stmt.psimps(2)[OF 2(1)] g_def st'_def st''_def by simp + next + case s3: (Some v') + with 2(1) `\ gas st \ g` n Pair KValue Value n2 p1 LStackloc v2 s3 + have "stmt (ASSIGN lv ex) env cd st = Normal ((), st''\gas := g'', stack := updateStore l (KValue v') (stack st)\)" + using stmt.psimps(2) g_def st'_def st''_def by simp + with stmt_def have "st6'= st''\gas := g'', stack := updateStore l (KValue v') (stack st)\" by simp + moreover from lexp_gas `\ gas st \ g` n2 p1 have "gas (st''\gas := g'', stack := updateStore l (KValue v') (stack st)\) \ gas (st'\gas := g'\)" using g_def st'_def by simp + moreover from msel_ssel_expr_load_rexp_gas(3)[of ex env cd st' "gas st - g"] `\ gas st \ g` n Pair KValue Value have "gas (st'\gas := g'\) \ gas (st\gas := gas st - g\)" using g_def by simp + ultimately show ?thesis by simp + qed + next + case (Calldata x2) + with 2(1) stmt_def `\ gas st \ g` n Pair KValue Value n2 p1 LStackloc show ?thesis using stmt.psimps(2) g_def st'_def st''_def by simp + next + case (Memory x3) + with 2(1) stmt_def `\ gas st \ g` n Pair KValue Value n2 p1 LStackloc show ?thesis using stmt.psimps(2) g_def st'_def st''_def by simp + next + case (Storage x4) + with 2(1) stmt_def `\ gas st \ g` n Pair KValue Value n2 p1 LStackloc show ?thesis using stmt.psimps(2) g_def st'_def st''_def by simp + qed + next + case (LMemloc l) + then show ?thesis + proof (cases b) + case v2: (Value t') + with 2(1) stmt_def `\ gas st \ g` n Pair KValue Value n2 p1 LMemloc show ?thesis using stmt.psimps(2) g_def st'_def st''_def by simp + next + case (Calldata x2) + with 2(1) stmt_def `\ gas st \ g` n Pair KValue Value n2 p1 LMemloc show ?thesis using stmt.psimps(2) g_def st'_def st''_def by simp + next + case (Memory x3) + then show ?thesis + proof (cases x3) + case (MTArray x11 x12) + with 2(1) stmt_def `\ gas st \ g` n Pair KValue Value n2 p1 LMemloc Memory show ?thesis using stmt.psimps(2) g_def st'_def st''_def by simp + next + case (MTValue t') + then show ?thesis + proof (cases "convert t t' v ") + case None + with 2(1) stmt_def `\ gas st \ g` n Pair KValue Value n2 p1 LMemloc Memory MTValue show ?thesis using stmt.psimps(2) g_def st'_def st''_def by simp + next + case s3: (Some v') + with 2(1) `\ gas st \ g` n Pair KValue Value n2 p1 LMemloc Memory MTValue s3 + have "stmt (ASSIGN lv ex) env cd st = Normal ((), st''\gas := g'', memory := updateStore l (MValue v') (memory st'')\)" + using stmt.psimps(2) g_def st'_def st''_def by simp + with stmt_def have "st6'= (st''\gas := g'', memory := updateStore l (MValue v') (memory st'')\)" by simp + moreover from lexp_gas `\ gas st \ g` n2 p1 have "gas (st''\gas := g'', stack := updateStore l (KValue v') (stack st)\) \ gas (st'\gas := g'\)" using g_def st'_def by simp + moreover from msel_ssel_expr_load_rexp_gas(3)[of ex env cd st' "gas st - g"] `\ gas st \ g` n Pair KValue Value n2 p1 have "gas (st'\gas := g'\) \ gas (st\gas := gas st - g\)" using g_def by simp + ultimately show ?thesis by simp + qed + qed + next + case (Storage x4) + with 2(1) stmt_def `\ gas st \ g` n Pair KValue Value n2 p1 LMemloc show ?thesis using stmt.psimps(2) g_def st'_def st''_def by simp + qed + next + case (LStoreloc l) + then show ?thesis + proof (cases b) + case v2: (Value t') + with 2(1) stmt_def `\ gas st \ g` n Pair KValue Value n2 p1 LStoreloc show ?thesis using stmt.psimps(2) g_def st'_def st''_def by simp + next + case (Calldata x2) + with 2(1) stmt_def `\ gas st \ g` n Pair KValue Value n2 p1 LStoreloc show ?thesis using stmt.psimps(2) g_def st'_def st''_def by simp + next + case (Memory x3) + with 2(1) stmt_def `\ gas st \ g` n Pair KValue Value n2 p1 LStoreloc show ?thesis using stmt.psimps(2) g_def st'_def st''_def by simp + next + case (Storage x4) + then show ?thesis + proof (cases x4) + case (STArray x11 x12) + with 2(1) stmt_def `\ gas st \ g` n Pair KValue Value n2 p1 LStoreloc Storage show ?thesis using stmt.psimps(2) g_def st'_def st''_def by simp + next + case (STMap x21 x22) + with 2(1) stmt_def `\ gas st \ g` n Pair KValue Value n2 p1 LStoreloc Storage show ?thesis using stmt.psimps(2) g_def st'_def st''_def by simp + next + case (STValue t') + then show ?thesis + proof (cases "convert t t' v ") + case None + with 2(1) stmt_def `\ gas st \ g` n Pair KValue Value n2 p1 LStoreloc Storage STValue show ?thesis using stmt.psimps(2) g_def st'_def st''_def by simp + next + case s3: (Some v') + with 2(1) `\ gas st \ g` n Pair KValue Value n2 p1 LStoreloc Storage STValue s3 + have "stmt (ASSIGN lv ex) env cd st = Normal ((), st'' \gas := g'', storage := (storage st'') (address env := fmupd l v' (storage st'' (address env)))\)" + using stmt.psimps(2) g_def st'_def st''_def by simp + with stmt_def have "st6'= st'' \gas := g'', storage := (storage st'') (address env := fmupd l v' (storage st'' (address env)))\" by simp + moreover from lexp_gas `\ gas st \ g` n Pair KValue Value n2 p1 have "gas (st''\gas := g'', stack := updateStore l (KValue v') (stack st)\) \ gas (st'\gas := g'\)" using g_def by simp + moreover from msel_ssel_expr_load_rexp_gas(3)[of ex env cd st' "gas st - g"] `\ gas st \ g` n Pair KValue Value n2 p1 have "gas (st'\gas := g'\) \ gas (st\gas := gas st - g\)" using g_def by simp + ultimately show ?thesis by simp + qed + qed + qed + qed + qed + next + case (e x) + with 2(1) stmt_def `\ gas st \ g` n Pair KValue Value show ?thesis using stmt.psimps(2) g_def st'_def st''_def by simp + qed + next + case (Calldata x2) + with 2(1) stmt_def `\ gas st \ g` n Pair KValue show ?thesis using stmt.psimps(2) g_def st'_def by simp + next + case (Memory x3) + with 2(1) stmt_def `\ gas st \ g` n Pair KValue show ?thesis using stmt.psimps(2) g_def st'_def by simp + next + case (Storage x4) + with 2(1) stmt_def `\ gas st \ g` n Pair KValue show ?thesis using stmt.psimps(2) g_def st'_def by simp + qed + next + case (KCDptr p) + then show ?thesis + proof (cases c) + case (Value x1) + with 2(1) stmt_def `\ gas st \ g` n Pair KCDptr show ?thesis using stmt.psimps(2) g_def st'_def by simp + next + case (Calldata x2) + then show ?thesis + proof (cases x2) + case (MTArray x t) + then show ?thesis + proof (cases "lexp lv env cd st'' g'") + case n2: (n a g'') + define st''' where "st''' = st''\gas := g''\" + then show ?thesis + proof (cases a) + case p2: (Pair a b) + then show ?thesis + proof (cases a) + case (LStackloc l) + then show ?thesis + proof (cases b) + case v2: (Value t') + with 2(1) stmt_def `\ gas st \ g` n Pair KCDptr Calldata MTArray n2 p2 LStackloc show ?thesis using stmt.psimps(2) g_def st'_def st''_def by simp + next + case c2: (Calldata x2) + with 2(1) stmt_def `\ gas st \ g` n Pair KCDptr Calldata MTArray n2 p2 LStackloc show ?thesis using stmt.psimps(2) g_def st'_def st''_def by simp + next + case (Memory x3) + then show ?thesis + proof (cases "accessStore l (stack st''')") + case None + with 2(1) stmt_def `\ gas st \ g` n Pair KCDptr Calldata MTArray n2 p2 LStackloc Memory show ?thesis using stmt.psimps(2) g_def st'_def st''_def st'''_def by simp + next + case s3: (Some a) + then show ?thesis + proof (cases a) + case (KValue x1) + with 2(1) stmt_def `\ gas st \ g` n Pair KCDptr Calldata MTArray n2 p2 LStackloc Memory s3 show ?thesis using stmt.psimps(2) g_def st'_def st''_def st'''_def by simp + next + case c3: (KCDptr x2) + with 2(1) stmt_def `\ gas st \ g` n Pair KCDptr Calldata MTArray n2 p2 LStackloc Memory s3 show ?thesis using stmt.psimps(2) g_def st'_def st''_def st'''_def by simp + next + case (KMemptr p') + then show ?thesis + proof (cases "cpm2m p p' x t cd (memory st''')") + case None + with 2(1) stmt_def `\ gas st \ g` n Pair KCDptr Calldata MTArray n2 p2 LStackloc Memory s3 KMemptr show ?thesis using stmt.psimps(2) g_def st'_def st''_def st'''_def by (simp split:if_split_asm) + next + case (Some m') + with `\ gas st \ g` n Pair KCDptr Calldata MTArray n2 p2 LStackloc Memory s3 KMemptr + have "stmt (ASSIGN lv ex) env cd st = Normal ((), st''' \memory := m'\)" + using stmt.psimps(2)[OF 2(1)] g_def st'_def st''_def st'''_def by simp + with stmt_def have "st6'= st''' \memory := m'\" by simp + moreover from lexp_gas `\ gas st \ g` n Pair KCDptr Calldata MTArray n2 p2 have "gas (st'''\memory := m'\) \ gas st''" using st''_def st'''_def by simp + moreover from msel_ssel_expr_load_rexp_gas(3)[of ex env cd st' "gas st - g"] `\ gas st \ g` n Pair have "gas st'' \ gas st'" using st'_def st''_def by simp + ultimately show ?thesis using st'_def by simp + qed + next + case (KStoptr p') + with 2(1) stmt_def `\ gas st \ g` n Pair KCDptr Calldata MTArray n2 p2 LStackloc Memory s3 show ?thesis using stmt.psimps(2) g_def st'_def st''_def st'''_def by simp + qed + qed + next + case (Storage x4) + then show ?thesis + proof (cases "accessStore l (stack st'')") + case None + with 2(1) stmt_def `\ gas st \ g` n Pair KCDptr Calldata MTArray n2 p2 LStackloc Storage show ?thesis using stmt.psimps(2) g_def st'_def st''_def st'''_def by simp + next + case s3: (Some a) + then show ?thesis + proof (cases a) + case (KValue x1) + with 2(1) stmt_def `\ gas st \ g` n Pair KCDptr Calldata MTArray n2 p2 LStackloc Storage s3 show ?thesis using stmt.psimps(2) g_def st'_def st''_def st'''_def by simp + next + case c3: (KCDptr x2) + with 2(1) stmt_def `\ gas st \ g` n Pair KCDptr Calldata MTArray n2 p2 LStackloc Storage s3 show ?thesis using stmt.psimps(2) g_def st'_def st''_def st'''_def by simp + next + case (KMemptr x3) + with 2(1) stmt_def `\ gas st \ g` n Pair KCDptr Calldata MTArray n2 p2 LStackloc Storage s3 show ?thesis using stmt.psimps(2) g_def st'_def st''_def st'''_def by simp + next + case (KStoptr p') + then show ?thesis + proof (cases "cpm2s p p' x t cd (storage st'' (address env))") + case None + with 2(1) stmt_def `\ gas st \ g` n Pair KCDptr Calldata MTArray n2 p2 LStackloc Storage s3 KStoptr show ?thesis using stmt.psimps(2) g_def st'_def st''_def st'''_def by simp + next + case (Some s') + with 2(1) `\ gas st \ g` n Pair KCDptr Calldata MTArray n2 p2 LStackloc Storage s3 KStoptr + have "stmt (ASSIGN lv ex) env cd st = Normal ((), st''' \storage := (storage st'') (address env := s')\)" + using stmt.psimps(2) g_def st'_def st''_def st'''_def by simp + with stmt_def have "st6'= st''' \storage := (storage st'') (address env := s')\" by simp + moreover from lexp_gas `\ gas st \ g` n Pair KCDptr Calldata MTArray n2 p2 have "gas st''' \ gas st''" using g_def st''_def st'''_def by simp + moreover from msel_ssel_expr_load_rexp_gas(3)[of ex env cd st' "gas st - g"] `\ gas st \ g` n Pair have "gas st'' \ gas st'" using g_def st'_def st''_def by simp + ultimately show ?thesis using st'_def by simp + qed + qed + qed + qed + next + case (LMemloc l) + then show ?thesis + proof (cases "cpm2m p l x t cd (memory st''')") + case None + with 2(1) stmt_def `\ gas st \ g` n Pair KCDptr Calldata MTArray n2 p2 LMemloc show ?thesis using stmt.psimps(2) g_def st'_def st''_def st'''_def by (simp split:if_split_asm) + next + case (Some m) + with `\ gas st \ g` n Pair KCDptr Calldata MTArray n2 p2 LMemloc + have "stmt (ASSIGN lv ex) env cd st = Normal ((), st'''\memory := m\)" + using stmt.psimps(2)[OF 2(1)] g_def st'_def st''_def st'''_def by simp + with stmt_def have "st6'= (st'''\memory := m\)" by simp + moreover from lexp_gas `\ gas st \ g` n Pair KCDptr Calldata MTArray n2 p2 have "gas st''' \ gas st''" using st''_def st'''_def by simp + moreover from msel_ssel_expr_load_rexp_gas(3)[of ex env cd st' "gas st - g"] `\ gas st \ g` n Pair have "gas st'' \ gas st'" using st'_def st''_def by simp + ultimately show ?thesis using st'_def by simp + qed + next + case (LStoreloc l) + then show ?thesis + proof (cases "cpm2s p l x t cd (storage st'' (address env))") + case None + with 2(1) stmt_def `\ gas st \ g` n Pair KCDptr Calldata MTArray n2 p2 LStoreloc show ?thesis using stmt.psimps(2) g_def st'_def st''_def st'''_def by simp + next + case (Some s) + with `\ gas st \ g` n Pair KCDptr Calldata MTArray n2 p2 LStoreloc + have "stmt (ASSIGN lv ex) env cd st = Normal ((), st''' \storage := (storage st'') (address env := s)\)" + using stmt.psimps(2)[OF 2(1)] g_def st'_def st''_def st'''_def by simp + with stmt_def have "st6'= (st'''\storage := (storage st'') (address env := s)\)" by simp + moreover from lexp_gas `\ gas st \ g` n Pair KCDptr Calldata MTArray n2 p2 have "gas st''' \ gas st''" using st''_def st'''_def by simp + moreover from msel_ssel_expr_load_rexp_gas(3)[of ex env cd st' "gas st - g"] `\ gas st \ g` n Pair have "gas st'' \ gas st'" using st'_def st''_def by simp + ultimately show ?thesis using st'_def by simp + qed + qed + qed + next + case (e x) + with 2(1) stmt_def `\ gas st \ g` n Pair KCDptr Calldata MTArray show ?thesis using stmt.psimps(2) g_def st'_def st''_def by simp + qed + next + case (MTValue x2) + with 2(1) stmt_def `\ gas st \ g` n Pair KCDptr Calldata show ?thesis using stmt.psimps(2) g_def st'_def st''_def by simp + qed + next + case (Memory x3) + with 2(1) stmt_def `\ gas st \ g` n Pair KCDptr show ?thesis using stmt.psimps(2) g_def st'_def st''_def by simp + next + case (Storage x4) + with 2(1) stmt_def `\ gas st \ g` n Pair KCDptr show ?thesis using stmt.psimps(2) g_def st'_def st''_def by simp + qed + next + case (KMemptr p) + then show ?thesis + proof (cases c) + case (Value x1) + with 2(1) stmt_def `\ gas st \ g` n Pair KMemptr show ?thesis using stmt.psimps(2) g_def st'_def st''_def by simp + next + case (Calldata x2) + with 2(1) stmt_def `\ gas st \ g` n Pair KMemptr show ?thesis using stmt.psimps(2) g_def st'_def st''_def by simp + next + case (Memory x3) + then show ?thesis + proof (cases x3) + case (MTArray x t) + then show ?thesis + proof (cases "lexp lv env cd st'' g'") + case n2: (n a g'') + define st''' where "st''' = st''\gas := g''\" + then show ?thesis + proof (cases a) + case p2: (Pair a b) + then show ?thesis + proof (cases a) + case (LStackloc l) + then show ?thesis + proof (cases b) + case v2: (Value t') + with 2(1) stmt_def `\ gas st \ g` n Pair KMemptr Memory MTArray n2 p2 LStackloc show ?thesis using stmt.psimps(2) g_def st'_def st''_def by simp + next + case c2: (Calldata x2) + with 2(1) stmt_def `\ gas st \ g` n Pair KMemptr Memory MTArray n2 p2 LStackloc show ?thesis using stmt.psimps(2) g_def st'_def st''_def by simp + next + case m2: (Memory x3) + with 2(1) `\ gas st \ g` n Pair KMemptr Memory MTArray n2 p2 LStackloc + have "stmt (ASSIGN lv ex) env cd st = Normal ((), st'''\stack := updateStore l (KMemptr p) (stack st'')\)" + using stmt.psimps(2)[OF 2(1)] g_def st'_def st''_def st'''_def by simp + with stmt_def have "st6'= st'''\stack := updateStore l (KMemptr p) (stack st'')\" by simp + moreover from lexp_gas `\ gas st \ g` n Pair KMemptr Memory MTArray n2 p2 have "gas st''' \ gas st''" using st''_def st'''_def by simp + moreover from msel_ssel_expr_load_rexp_gas(3)[of ex env cd st' "gas st - g"] `\ gas st \ g` n Pair have "gas st'' \ gas st'" using st'_def st''_def by simp + ultimately show ?thesis using st'_def by simp + next + case (Storage x4) + then show ?thesis + proof (cases "accessStore l (stack st''')") + case None + with 2(1) stmt_def `\ gas st \ g` n Pair KMemptr Memory MTArray n2 p2 LStackloc Storage show ?thesis using stmt.psimps(2) g_def st'_def st''_def st'''_def by simp + next + case s3: (Some a) + then show ?thesis + proof (cases a) + case (KValue x1) + with 2(1) stmt_def `\ gas st \ g` n Pair KMemptr Memory MTArray n2 p2 LStackloc Storage s3 show ?thesis using stmt.psimps(2) g_def st'_def st''_def st'''_def by simp + next + case c3: (KCDptr x2) + with 2(1) stmt_def `\ gas st \ g` n Pair KMemptr Memory MTArray n2 p2 LStackloc Storage s3 show ?thesis using stmt.psimps(2) g_def st'_def st''_def st'''_def by simp + next + case m3: (KMemptr x3) + with 2(1) stmt_def `\ gas st \ g` n Pair KMemptr Memory MTArray n2 p2 LStackloc Storage s3 show ?thesis using stmt.psimps(2) g_def st'_def st''_def st'''_def by simp + next + case (KStoptr p') + then show ?thesis + proof (cases "cpm2s p p' x t (memory st''') (storage st''' (address env))") + case None + with 2(1) stmt_def `\ gas st \ g` n Pair KMemptr Memory MTArray n2 p2 LStackloc Storage s3 KStoptr show ?thesis using stmt.psimps(2) g_def st'_def st''_def st'''_def by simp + next + case (Some s) + with 2(1) `\ gas st \ g` n Pair KMemptr Memory MTArray n2 p2 LStackloc Storage s3 KStoptr + have "stmt (ASSIGN lv ex) env cd st = Normal ((), st'''\storage := (storage st''') (address env := s)\)" + using stmt.psimps(2) g_def st'_def st''_def st'''_def by simp + with stmt_def have "st6'= st'''\storage := (storage st''') (address env := s)\" by simp + moreover from lexp_gas `\ gas st \ g` n Pair KMemptr Memory MTArray n2 p2 have "gas st''' \ gas st''" using g_def st'_def st''_def st'''_def by simp + moreover from msel_ssel_expr_load_rexp_gas(3)[of ex env cd st' "gas st - g"] `\ gas st \ g` n Pair have "gas st'' \ gas st'" using st'_def st''_def by simp + ultimately show ?thesis using st'_def by simp + qed + qed + qed + qed + next + case (LMemloc l) + with 2(1) `\ gas st \ g` n Pair KMemptr Memory MTArray n2 p2 LMemloc + have "stmt (ASSIGN lv ex) env cd st = Normal ((), st'''\memory := updateStore l (MPointer p) (memory st''')\)" + using stmt.psimps(2) g_def st'_def st''_def st'''_def by simp + with stmt_def have "st6'= st'''\memory := updateStore l (MPointer p) (memory st''')\" by simp + moreover from lexp_gas `\ gas st \ g` n Pair KMemptr Memory MTArray n2 p2 have "gas st''' \ gas st''" using g_def st'_def st''_def st'''_def by simp + moreover from msel_ssel_expr_load_rexp_gas(3)[of ex env cd st' "gas st - g"] `\ gas st \ g` n Pair have "gas st'' \ gas st'" using st'_def st''_def by simp + ultimately show ?thesis using st'_def by simp + next + case (LStoreloc l) + then show ?thesis + proof (cases "cpm2s p l x t (memory st''') (storage st'' (address env))") + case None + with 2(1) stmt_def `\ gas st \ g` n Pair KMemptr Memory MTArray n2 p2 LStoreloc show ?thesis using stmt.psimps(2) g_def using st'_def st''_def st'''_def by simp + next + case (Some s) + with 2(1) `\ gas st \ g` n Pair KMemptr Memory MTArray n2 p2 LStoreloc + have "stmt (ASSIGN lv ex) env cd st = Normal ((), st'''\storage := (storage st''') (address env := s)\)" + using stmt.psimps(2) g_def st'_def st''_def st'''_def by simp + with stmt_def have "st6'= st'''\storage := (storage st''') (address env := s)\" by simp + moreover from lexp_gas `\ gas st \ g` n Pair KMemptr Memory MTArray n2 p2 have "gas st''' \ gas st''" using g_def st'_def st''_def st'''_def by simp + moreover from msel_ssel_expr_load_rexp_gas(3)[of ex env cd st' "gas st - g"] `\ gas st \ g` n Pair have "gas st'' \ gas st'" using st'_def st''_def by simp + ultimately show ?thesis using st'_def by simp + qed + qed + qed + next + case (e _) + with 2(1) stmt_def `\ gas st \ g` n Pair KMemptr Memory MTArray show ?thesis using stmt.psimps(2) g_def st'_def st''_def by simp + qed + next + case (MTValue _) + with 2(1) stmt_def `\ gas st \ g` n Pair KMemptr Memory show ?thesis using stmt.psimps(2) g_def st'_def st''_def by simp + qed + next + case (Storage x4) + with 2(1) stmt_def `\ gas st \ g` n Pair KMemptr show ?thesis using stmt.psimps(2) g_def st'_def st''_def by simp + qed + next + case (KStoptr p) + then show ?thesis + proof (cases c) + case (Value x1) + with 2(1) stmt_def `\ gas st \ g` n Pair KStoptr show ?thesis using stmt.psimps(2) g_def st'_def st''_def by simp + next + case (Calldata x2) + with 2(1) stmt_def `\ gas st \ g` n Pair KStoptr show ?thesis using stmt.psimps(2) g_def st'_def st''_def by simp + next + case (Memory x3) + with 2(1) stmt_def `\ gas st \ g` n Pair KStoptr show ?thesis using stmt.psimps(2) g_def st'_def st''_def by simp + next + case (Storage x4) + then show ?thesis + proof (cases x4) + case (STArray x t) + then show ?thesis + proof (cases "lexp lv env cd st'' g'") + case n2: (n a g'') + define st''' where "st''' = st''\gas := g''\" + then show ?thesis + proof (cases a) + case p2: (Pair a b) + then show ?thesis + proof (cases a) + case (LStackloc l) + then show ?thesis + proof (cases b) + case v2: (Value t') + with 2(1) stmt_def `\ gas st \ g` n Pair KStoptr Storage STArray n2 p2 LStackloc show ?thesis using stmt.psimps(2) g_def st'_def st''_def by simp + next + case c2: (Calldata x2) + with 2(1) stmt_def `\ gas st \ g` n Pair KStoptr Storage STArray n2 p2 LStackloc show ?thesis using stmt.psimps(2) g_def st'_def st''_def by simp + next + case (Memory x3) + then show ?thesis + proof (cases "accessStore l (stack st''')") + case None + with 2(1) stmt_def `\ gas st \ g` n Pair KStoptr Storage STArray n2 p2 LStackloc Memory show ?thesis using stmt.psimps(2) g_def st'_def st''_def st'''_def by simp + next + case s3: (Some a) + then show ?thesis + proof (cases a) + case (KValue x1) + with 2(1) stmt_def `\ gas st \ g` n Pair KStoptr Storage STArray n2 p2 LStackloc Memory s3 show ?thesis using stmt.psimps(2) g_def st'_def st''_def st'''_def by simp + next + case c3: (KCDptr x2) + with 2(1) stmt_def `\ gas st \ g` n Pair KStoptr Storage STArray n2 p2 LStackloc Memory s3 show ?thesis using stmt.psimps(2) g_def st'_def st''_def st'''_def by simp + next + case (KMemptr p') + then show ?thesis + proof (cases "cps2m p p' x t (storage st''' (address env)) (memory st''')") + case None + with 2(1) stmt_def `\ gas st \ g` n Pair KStoptr Storage STArray n2 p2 LStackloc Memory s3 KMemptr show ?thesis using stmt.psimps(2) g_def st'_def st''_def st'''_def by simp + next + case (Some m) + with 2(1) `\ gas st \ g` n Pair KStoptr Storage STArray n2 p2 LStackloc Memory s3 KMemptr + have "stmt (ASSIGN lv ex) env cd st = Normal ((), st'''\memory := m\)" + using stmt.psimps(2) g_def st'_def st''_def st'''_def by simp + with stmt_def have "st6'= st'''\memory := m\" by simp + moreover from lexp_gas `\ gas st \ g` n Pair KMemptr Storage STArray n2 p2 have "gas (st'''\memory := m\) \ gas st''" using g_def st'_def st''_def st'''_def by simp + moreover from msel_ssel_expr_load_rexp_gas(3)[of ex env cd st' "gas st - g"] `\ gas st \ g` n Pair have "gas st'' \ gas st'" using st'_def st''_def by simp + ultimately show ?thesis using st'_def by simp + qed + next + case sp2: (KStoptr p') + with 2(1) stmt_def `\ gas st \ g` n Pair KStoptr Storage STArray n2 p2 LStackloc Memory s3 show ?thesis using stmt.psimps(2) g_def st'_def st''_def st'''_def by simp + qed + qed + next + case st2: (Storage x4) + with 2(1) `\ gas st \ g` n Pair KStoptr Storage STArray n2 p2 LStackloc + have "stmt (ASSIGN lv ex) env cd st = Normal ((), st'''\stack := updateStore l (KStoptr p) (stack st''')\)" + using stmt.psimps(2) g_def st'_def st''_def st'''_def by simp + with stmt_def have "st6'= st'''\stack := updateStore l (KStoptr p) (stack st''')\" by simp + moreover from lexp_gas `\ gas st \ g` n Pair KStoptr Storage STArray n2 p2 have "gas (st'''\stack := updateStore l (KStoptr p) (stack st''')\) \ gas st''" using g_def st'_def st''_def st'''_def by simp + moreover from msel_ssel_expr_load_rexp_gas(3)[of ex env cd st' "gas st - g"] `\ gas st \ g` n Pair have "gas st'' \ gas st'" using st'_def st''_def by simp + ultimately show ?thesis using st'_def by simp + qed + next + case (LMemloc l) + then show ?thesis + proof (cases "cps2m p l x t (storage st'' (address env)) (memory st'')") + case None + with 2(1) stmt_def `\ gas st \ g` n Pair KStoptr Storage STArray n2 p2 LMemloc show ?thesis using stmt.psimps(2) g_def st'_def st''_def st'''_def by simp + next + case (Some m) + with 2(1) `\ gas st \ g` n Pair KStoptr Storage STArray n2 p2 LMemloc + have "stmt (ASSIGN lv ex) env cd st = Normal ((), st'''\memory := m\)" + using stmt.psimps(2) g_def st'_def st''_def st'''_def by simp + with stmt_def have "st6'= (st'''\memory := m\)" by simp + moreover from lexp_gas `\ gas st \ g` n Pair KStoptr Storage STArray n2 p2 have "gas (st'''\memory := m\) \ gas st''" using g_def st'_def st''_def st'''_def by simp + moreover from msel_ssel_expr_load_rexp_gas(3)[of ex env cd st' "gas st - g"] `\ gas st \ g` n Pair have "gas st'' \ gas st'" using st'_def st''_def by simp + ultimately show ?thesis using st'_def by simp + qed + next + case (LStoreloc l) + then show ?thesis + proof (cases "copy p l x t (storage st'' (address env))") + case None + with 2(1) stmt_def `\ gas st \ g` n Pair KStoptr Storage STArray n2 p2 LStoreloc show ?thesis using stmt.psimps(2) g_def st'_def st''_def st'''_def by simp + next + case (Some s) + with 2(1) `\ gas st \ g` n Pair KStoptr Storage STArray n2 p2 LStoreloc + have "stmt (ASSIGN lv ex) env cd st = Normal ((), st'''\storage := (storage st''') (address env := s)\)" + using stmt.psimps(2) g_def st'_def st''_def st'''_def by simp + with stmt_def have "st6'= st'''\storage := (storage st''') (address env := s)\" by simp + moreover from lexp_gas `\ gas st \ g` n Pair KStoptr Storage STArray n2 p2 have "gas (st'''\storage := (storage st''') (address env := s')\) \ gas st''" using g_def st'_def st''_def st'''_def by simp + moreover from msel_ssel_expr_load_rexp_gas(3)[of ex env cd st' "gas st - g"] `\ gas st \ g` n Pair have "gas st'' \ gas st'" using st'_def st''_def by simp + ultimately show ?thesis using st'_def by simp + qed + qed + qed + next + case (e x) + with 2(1) stmt_def `\ gas st \ g` n Pair KStoptr Storage STArray show ?thesis using stmt.psimps(2) g_def st'_def st''_def by simp + qed + next + case (STMap t t') + then show ?thesis + proof (cases "lexp lv env cd st'' g'") + case n2: (n a g'') + define st''' where "st''' = st''\gas := g''\" + then show ?thesis + proof (cases a) + case p2: (Pair a b) + then show ?thesis + proof (cases a) + case (LStackloc l) + with 2(1) `\ gas st \ g` n Pair KStoptr Storage STMap n2 p2 + have "stmt (ASSIGN lv ex) env cd st = Normal ((), st''' \stack := updateStore l (KStoptr p) (stack st''')\)" + using stmt.psimps(2) g_def st'_def st''_def st'''_def by simp + with stmt_def have "st6'= st'''\stack := updateStore l (KStoptr p) (stack st''')\" by simp + moreover from lexp_gas `\ gas st \ g` n Pair KStoptr Storage STMap n2 p2 have "gas (st'''\stack := updateStore l (KStoptr p) (stack st''')\) \ gas st''" using g_def st'_def st''_def st'''_def by simp + moreover from msel_ssel_expr_load_rexp_gas(3)[of ex env cd st' "gas st - g"] `\ gas st \ g` n Pair have "gas st'' \ gas st'" using st'_def st''_def by simp + ultimately show ?thesis using st'_def by simp + next + case (LMemloc x2) + with 2(1) stmt_def `\ gas st \ g` n Pair KStoptr Storage STMap n2 p2 show ?thesis using stmt.psimps(2) g_def st'_def st''_def by simp + next + case (LStoreloc x3) + with 2(1) stmt_def `\ gas st \ g` n Pair KStoptr Storage STMap n2 p2 show ?thesis using stmt.psimps(2) g_def st'_def st''_def by simp + qed + qed + next + case (e x) + with 2(1) stmt_def `\ gas st \ g` n Pair KStoptr Storage STMap show ?thesis using stmt.psimps(2) g_def st'_def st''_def by simp + qed + next + case (STValue x3) + with 2(1) stmt_def `\ gas st \ g` n Pair KStoptr Storage show ?thesis using stmt.psimps(2) g_def st'_def st''_def by simp + qed + qed + qed + qed + next + case (e x) + with 2(1) stmt_def `\ gas st \ g` show ?thesis using stmt.psimps(2) g_def st'_def by simp + qed + qed + qed +next + case (3 s1 s2 e cd st) + define g where "g = costs (COMP s1 s2) e cd st" + show ?case + proof (rule allI[OF impI]) + fix st6' + assume stmt_def: "stmt (COMP s1 s2) e cd st = Normal ((), st6')" + then show "gas st6' \ gas st" + proof cases + assume "gas st \ g" + with 3(1) stmt_def g_def show ?thesis using stmt.psimps(3) by simp + next + assume "\ gas st \ g" + show ?thesis + proof (cases "stmt s1 e cd (st\gas := gas st - g\)") + case (n a st') + with 3(1) stmt_def `\ gas st \ g` have "stmt (COMP s1 s2) e cd st = stmt s2 e cd st'" using stmt.psimps(3)[of s1 s2 e cd st] g_def by (simp add:Let_def split:unit.split_asm) + with 3(3) stmt_def \\ gas st \ g\ n have "gas st6' \ gas st'" using g_def by simp + moreover from 3(2)[where ?s'a="st\gas := gas st - g\"] \\ gas st \ g\ n have "gas st' \ gas st" using g_def by simp + ultimately show ?thesis by simp + next + case (e x) + with 3 stmt_def `\ gas st \ g` show ?thesis using stmt.psimps(3)[of s1 s2 e cd st] g_def by (simp split: Ex.split_asm) + qed + qed + qed +next + case (4 ex s1 s2 e cd st) + define g where "g = costs (ITE ex s1 s2) e cd st" + show ?case + proof (rule allI[OF impI]) + fix st6' + assume stmt_def: "stmt (ITE ex s1 s2) e cd st = Normal ((), st6')" + then show "gas st6' \ gas st" + proof cases + assume "gas st \ g" + with 4(1) stmt_def show ?thesis using stmt.psimps(4) g_def by simp + next + assume "\ gas st \ g" + then have l1: "assert Gas (\st. costs (ITE ex s1 s2) e cd st < gas st) st = Normal ((), st) " using g_def by simp + define st' where "st' = st\gas := gas st - g\" + then have l2: " modify (\st. st\gas := gas st - costs (ITE ex s1 s2) e cd st\) st = Normal ((), st')" using g_def by simp + show ?thesis + proof (cases "expr ex e cd st' (gas st - g)") + case (n a g') + define st'' where "st'' = st'\gas := g'\" + with n have l3: "toState (expr ex e cd) st' = Normal (a, st'')" using st'_def by simp + then show ?thesis + proof (cases a) + case (Pair b c) + then show ?thesis + proof (cases b) + case (KValue b) + then show ?thesis + proof (cases c) + case (Value x1) + then show ?thesis + proof (cases x1) + case (TSInt x1) + with 4(1) stmt_def `\ gas st \ g` n Pair KValue Value show ?thesis using stmt.psimps(4) g_def st'_def by simp + next + case (TUInt x2) + with 4(1) stmt_def `\ gas st \ g` n Pair KValue Value show ?thesis using stmt.psimps(4) g_def st'_def by simp + next + case TBool + then show ?thesis + proof cases + assume "b = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True" + with 4(1) `\ gas st \ g` n Pair KValue Value TBool have "stmt (ITE ex s1 s2) e cd st = stmt s1 e cd st''" using stmt.psimps(4) g_def st'_def st''_def by simp + with 4(2)[OF l1 l2 l3] stmt_def `\ gas st \ g` n Pair KValue Value TBool `b = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True` have "gas st6' \ gas st''" using g_def by simp + moreover from msel_ssel_expr_load_rexp_gas(3)[of ex e cd st' "gas st - g"] `\ gas st \ g` n Pair KValue Value TBool have "gas st'' \ gas st'" using st'_def st''_def by simp + ultimately show ?thesis using st'_def by simp + next + assume nt: "\ b = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True" + show ?thesis + proof cases + assume "b = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l False" + with 4(1) `\ gas st \ g` n Pair KValue Value TBool nt have "stmt (ITE ex s1 s2) e cd st = stmt s2 e cd st''" using stmt.psimps(4) g_def st'_def st''_def by simp + with 4(3)[OF l1 l2 l3] stmt_def `\ gas st \ g` n Pair KValue Value TBool nt `b = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l False` have "gas st6' \ gas st''" using g_def by simp + moreover from msel_ssel_expr_load_rexp_gas(3)[of ex e cd st' "gas st - g"] `\ gas st \ g` n Pair KValue Value TBool have "gas st'' \ gas st'" using st'_def st''_def by simp + ultimately show ?thesis using st'_def by simp + next + assume "\ b = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l False" + with 4(1) stmt_def `\ gas st \ g` n Pair KValue Value TBool nt show ?thesis using stmt.psimps(4) g_def st'_def st''_def by simp + qed + qed + next + case TAddr + with 4(1) stmt_def `\ gas st \ g` n Pair KValue Value show ?thesis using stmt.psimps(4) g_def st'_def st''_def by simp + qed + next + case (Calldata x2) + with 4(1) stmt_def `\ gas st \ g` n Pair KValue show ?thesis using stmt.psimps(4) g_def st'_def st''_def by simp + next + case (Memory x3) + with 4(1) stmt_def `\ gas st \ g` n Pair KValue show ?thesis using stmt.psimps(4) g_def st'_def st''_def by simp + next + case (Storage x4) + with 4(1) stmt_def `\ gas st \ g` n Pair KValue show ?thesis using stmt.psimps(4) g_def st'_def st''_def by simp + qed + next + case (KCDptr x2) + with 4(1) stmt_def `\ gas st \ g` n Pair show ?thesis using stmt.psimps(4) g_def st'_def st''_def by simp + next + case (KMemptr x3) + with 4(1) stmt_def `\ gas st \ g` n Pair show ?thesis using stmt.psimps(4) g_def st'_def st''_def by simp + next + case (KStoptr x4) + with 4(1) stmt_def `\ gas st \ g` n Pair show ?thesis using stmt.psimps(4) g_def st'_def st''_def by simp + qed + qed + next + case (e e) + with 4(1) stmt_def `\ gas st \ g` show ?thesis using stmt.psimps(4) g_def st'_def by simp + qed + qed + qed +next + case (5 ex s0 e cd st) + define g where "g = costs (WHILE ex s0) e cd st" + show ?case + proof (rule allI[OF impI]) + fix st6' + assume stmt_def: "stmt (WHILE ex s0) e cd st = Normal ((), st6')" + then show "gas st6' \ gas st" + proof cases + assume "gas st \ g" + with 5(1) stmt_def show ?thesis using stmt.psimps(5) g_def by simp + next + assume gcost: "\ gas st \ g" + then have l1: "assert Gas (\st. costs (WHILE ex s0) e cd st < gas st) st = Normal ((), st) " using g_def by simp + define st' where "st' = st\gas := gas st - g\" + then have l2: " modify (\st. st\gas := gas st - costs (WHILE ex s0) e cd st\) st = Normal ((), st')" using g_def by simp + show ?thesis + proof (cases "expr ex e cd st' (gas st - g)") + case (n a g') + define st'' where "st'' = st'\gas := g'\" + with n have l3: "toState (expr ex e cd) st' = Normal (a, st'')" using st'_def by simp + then show ?thesis + proof (cases a) + case (Pair b c) + then show ?thesis + proof (cases b) + case (KValue b) + then show ?thesis + proof (cases c) + case (Value x1) + then show ?thesis + proof (cases x1) + case (TSInt x1) + with 5(1) stmt_def gcost n Pair KValue Value show ?thesis using stmt.psimps(5) g_def st'_def by simp + next + case (TUInt x2) + with 5(1) stmt_def gcost n Pair KValue Value show ?thesis using stmt.psimps(5) g_def st'_def by simp + next + case TBool + then show ?thesis + proof cases + assume "b = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True" + then show ?thesis + proof (cases "stmt s0 e cd st''") + case n2: (n a st''') + with 5(1) gcost n Pair KValue Value TBool `b = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True` have "stmt (WHILE ex s0) e cd st = stmt (WHILE ex s0) e cd st'''" using stmt.psimps(5)[of ex s0 e cd st] g_def st'_def st''_def by (simp add: Let_def split:unit.split_asm) + with 5(3) stmt_def gcost n2 Pair KValue Value TBool `b = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True` n have "gas st6' \ gas st'''" using g_def st'_def st''_def by simp + moreover from 5(2)[OF l1 l2 l3] gcost n2 Pair KValue Value TBool `b = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True` n have "gas st''' \ gas st''" using g_def by simp + moreover from msel_ssel_expr_load_rexp_gas(3)[of ex e cd st' "gas st - g"] `\ gas st \ g` n Pair KValue Value TBool have "gas st'' \ gas st'" using st'_def st''_def by simp + ultimately show ?thesis using st'_def by simp + next + case (e x) + with 5(1) stmt_def gcost n Pair KValue Value TBool `b = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True` show ?thesis using stmt.psimps(5) g_def st'_def st''_def by simp + qed + next + assume nt: "\ b = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True" + show ?thesis + proof cases + assume "b = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l False" + with 5(1) gcost n Pair KValue Value TBool nt have "stmt (WHILE ex s0) e cd st = return () st''" using stmt.psimps(5) g_def st'_def st''_def by simp + with stmt_def have "gas st6' \ gas st''" by simp + moreover from msel_ssel_expr_load_rexp_gas(3)[of ex e cd st' "gas st - g"] `\ gas st \ g` n Pair KValue Value TBool have "gas st'' \ gas st'" using st'_def st''_def by simp + ultimately show ?thesis using g_def st'_def st''_def by simp + next + assume "\ b = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l False" + with 5(1) stmt_def gcost n Pair KValue Value TBool nt show ?thesis using stmt.psimps(5) g_def st'_def st''_def by simp + qed + qed + next + case TAddr + with 5(1) stmt_def gcost n Pair KValue Value show ?thesis using stmt.psimps(5) g_def st'_def st''_def by simp + qed + next + case (Calldata x2) + with 5(1) stmt_def gcost n Pair KValue show ?thesis using stmt.psimps(5) g_def st'_def st''_def by simp + next + case (Memory x3) + with 5(1) stmt_def gcost n Pair KValue show ?thesis using stmt.psimps(5) g_def st'_def st''_def by simp + next + case (Storage x4) + with 5(1) stmt_def gcost n Pair KValue show ?thesis using stmt.psimps(5) g_def st'_def st''_def by simp + qed + next + case (KCDptr x2) + with 5(1) stmt_def gcost n Pair show ?thesis using stmt.psimps(5) g_def st'_def st''_def by simp + next + case (KMemptr x3) + with 5(1) stmt_def gcost n Pair show ?thesis using stmt.psimps(5) g_def st'_def st''_def by simp + next + case (KStoptr x4) + with 5(1) stmt_def gcost n Pair show ?thesis using stmt.psimps(5) g_def st'_def st''_def by simp + qed + qed + next + case (e e) + with 5(1) stmt_def gcost show ?thesis using stmt.psimps(5) g_def st'_def by simp + qed + qed + qed +next + case (6 i xe e cd st) + define g where "g = costs (INVOKE i xe) e cd st" + show ?case + proof (rule allI[OF impI]) + fix st6' assume a1: "stmt (INVOKE i xe) e cd st = Normal ((), st6')" + show "gas st6' \ gas st" + proof (cases) + assume "gas st \ g" + with 6(1) a1 show ?thesis using stmt.psimps(6) g_def by simp + next + assume gcost: "\ gas st \ g" + then have l1: "assert Gas (\st. costs (INVOKE i xe) e cd st < gas st) st = Normal ((), st) " using g_def by simp + define st' where "st' = st\gas := gas st - g\" + then have l2: "modify (\st. st\gas := gas st - costs (INVOKE i xe) e cd st\) st = Normal ((), st')" using g_def by simp + then show ?thesis + proof (cases "ep $$ contract e") + case None + with 6(1) a1 gcost show ?thesis using stmt.psimps(6) g_def by simp + next + case (Some x) + then have l3: "option Err (\_. ep $$ contract e) st' = Normal (x, st')" by simp + then show ?thesis + proof (cases x) + case (fields ct _ _) + then show ?thesis + proof (cases "fmlookup ct i") + case None + with 6(1) g_def a1 gcost Some fields show ?thesis using stmt.psimps(6) by simp + next + case s1: (Some a) + then show ?thesis + proof (cases a) + case (Method x1) + then show ?thesis + proof (cases x1) + case p1: (fields fp ext f) + then show ?thesis + proof (cases ext) + case True + with 6(1) a1 g_def gcost Some fields s1 Method p1 show ?thesis using stmt.psimps(6) st'_def by auto + next + case False + then have l4: "(case ct $$ i of None \ throw Err | Some (Method (fp, True, f)) \ throw Err + | Some (Method (fp, False, f)) \ return (fp, f) | Some _ \ throw Err) st' = Normal ((fp,f),st')" using s1 Method p1 by simp + define m\<^sub>o e' + where "m\<^sub>o = memory st'" + and "e' = ffold (init ct) (emptyEnv (address e) (contract e) (sender e) (svalue e)) (fmdom ct)" + then show ?thesis + proof (cases "load False fp xe e' emptyStore emptyStore m\<^sub>o e cd st' (gas st - g)") + case s4: (n a g') + define st'' where "st'' = st'\gas := g'\" + then show ?thesis + proof (cases a) + case f2: (fields e\<^sub>l cd\<^sub>l k\<^sub>l m\<^sub>l) + then have l5: "toState (load False fp xe e' emptyStore emptyStore m\<^sub>o e cd) st' = Normal ((e\<^sub>l, cd\<^sub>l, k\<^sub>l, m\<^sub>l), st'')" using st'_def st''_def s4 by simp + define k\<^sub>o where "k\<^sub>o = stack st'" + then show ?thesis + proof (cases "stmt f e\<^sub>l cd\<^sub>l (st''\stack:=k\<^sub>l, memory:=m\<^sub>l\)") + case n2: (n a st''') + with a1 g_def gcost Some fields s1 Method p1 m\<^sub>o_def e'_def s4 f2 k\<^sub>o_def False + have "stmt (INVOKE i xe) e cd st = Normal ((), st'''\stack:=k\<^sub>o\)" + using stmt.psimps(6)[OF 6(1)] st'_def st''_def by auto + with a1 have "gas st6' \ gas st'''" by auto + also from 6(2)[OF l1 l2 l3 fields l4 _ _ _ l5, where ?s'g="st''\stack := k\<^sub>l, memory := m\<^sub>l\"] n2 m\<^sub>o_def e'_def + have "\ \ gas st''" by auto + also from msel_ssel_expr_load_rexp_gas(4)[of False fp xe e' emptyStore emptyStore m\<^sub>o e cd st' "(gas st - g)"] have "\ \ gas st'" using s4 st'_def st''_def f2 by auto + finally show ?thesis using st'_def by simp + next + case (e x) + with 6(1) a1 g_def gcost Some fields s1 Method p1 m\<^sub>o_def e'_def s4 f2 show ?thesis using stmt.psimps(6) st'_def st''_def False by auto + qed + qed + next + case (e x) + with 6(1) a1 g_def gcost Some fields s1 Method p1 m\<^sub>o_def e'_def show ?thesis using stmt.psimps(6) st'_def False by auto + qed + qed + qed + next + case (Function _) + with 6(1) g_def a1 gcost Some fields s1 show ?thesis using stmt.psimps(6) by simp + next + case (Var _) + with 6(1) g_def a1 gcost Some fields s1 show ?thesis using stmt.psimps(6) by simp + qed + qed + qed + qed + qed + qed +next + case (7 ad i xe val e cd st) + define g where "g = costs (EXTERNAL ad i xe val) e cd st" + show ?case + proof (rule allI[OF impI]) + fix st6' assume a1: "stmt (EXTERNAL ad i xe val) e cd st = Normal ((), st6')" + show "gas st6' \ gas st" + proof (cases) + assume "gas st \ g" + with 7(1) a1 show ?thesis using stmt.psimps(7) g_def by simp + next + assume gcost: "\ gas st \ g" + then have l1: "assert Gas (\st. costs (EXTERNAL ad i xe val) e cd st < gas st) st = Normal ((), st) " using g_def by simp + define st' where "st' = st\gas := gas st - g\" + then have l2: " modify (\st. st\gas := gas st - costs (EXTERNAL ad i xe val) e cd st\) st = Normal ((), st')" using g_def by simp + then show ?thesis + proof (cases "expr ad e cd st' (gas st - g)") + case (n a0 g') + define st'' where "st'' = st'\gas := g'\" + with n have l3: "toState (expr ad e cd) st' = Normal (a0, st'')" using st'_def by simp + then show ?thesis + proof (cases a0) + case (Pair b c) + then show ?thesis + proof (cases b) + case (KValue adv) + then show ?thesis + proof (cases c) + case (Value x1) + then show ?thesis + proof (cases x1) + case (TSInt x1) + with 7(1) g_def a1 gcost n Pair KValue Value show ?thesis using stmt.psimps(7) st'_def by auto + next + case (TUInt x2) + with 7(1) g_def a1 gcost n Pair KValue Value show ?thesis using stmt.psimps(7) st'_def by auto + next + case TBool + with 7(1) g_def a1 gcost n Pair KValue Value show ?thesis using stmt.psimps(7) st'_def by auto + next + case TAddr + then have l4: "(case a0 of (KValue adv, Value TAddr) \ return adv | (KValue adv, Value _) \ throw Err | (KValue adv, _) \ throw Err + | (_, b) \ throw Err) st'' = Normal (adv, st'')" using Pair KValue Value by simp + then show ?thesis + proof (cases "adv = address e") + case True + with 7(1) g_def a1 gcost n Pair KValue Value TAddr show ?thesis using stmt.psimps(7) st'_def by auto + next + case False + then have l5: "assert Err (\_. adv \ address e) st'' = Normal ((), st'')" by simp + then show ?thesis + proof (cases "type (accounts st'' adv)") + case None + with 7(1) g_def a1 gcost n Pair KValue Value TAddr False show ?thesis using stmt.psimps(7) st'_def st''_def by auto + next + case (Some x2) + then show ?thesis + proof (cases x2) + case EOA + with 7(1) g_def a1 gcost n Pair KValue Value TAddr False Some show ?thesis using stmt.psimps(7) st'_def st''_def by auto + next + case (Contract c) + then have l6: "(\st. case type (accounts st adv) of Some (Contract c) \ return c st | _ \ throw Err st) st'' = Normal (c, st'')" using Some by (simp split:atype.split option.split) + then show ?thesis + proof (cases "ep $$ c") + case None + with 7(1) g_def a1 gcost n Pair KValue Value TAddr False Contract Some show ?thesis using stmt.psimps(7) st'_def st''_def by auto + next + case s2: (Some x) + then show ?thesis + proof (cases x) + case p2: (fields ct x0 fb) + then have l7: "option Err (\_. ep $$ c) st'' = Normal ((ct, x0, fb), st'')" using s2 by simp + then show ?thesis + proof (cases "expr val e cd st'' (gas st'')") + case n1: (n kv g'') + define st''' where "st''' = st''\gas := g''\" + with n1 have l8: "toState (expr val e cd) st'' = Normal (kv, st''')" by simp + then show ?thesis + proof (cases kv) + case p3: (Pair a b) + then show ?thesis + proof (cases a) + case k2: (KValue v) + then show ?thesis + proof (cases b) + case v: (Value t) + then have l9: "(case kv of (KValue v, Value t) \ return (v, t) | (KValue v, _) \ throw Err | (_, b) \ throw Err) st''' = Normal ((v,t), st''')" using n1 p3 k2 by simp + show ?thesis + proof (cases "convert t (TUInt 256) v") + case None + with 7(1) g_def a1 gcost n Pair KValue Value TAddr False Contract Some s2 p2 None n1 p3 k2 v False show ?thesis using stmt.psimps(7)[OF 7(1)] st'_def st''_def st'''_def by simp + next + case s3: (Some v') + define e' where "e' = ffold (init ct) (emptyEnv adv c (address e) v') (fmdom ct)" + show ?thesis + proof (cases "fmlookup ct i") + case None + show ?thesis + proof (cases "transfer (address e) adv v' (accounts st''')") + case n2: None + with 7(1) g_def a1 gcost n Pair KValue Value TAddr False Contract Some s2 p2 None n1 p3 k2 v False s3 show ?thesis using stmt.psimps(7)[OF 7(1)] st'_def st''_def st'''_def by simp + next + case s4: (Some acc) + then have l10: "option Err (\st. transfer (address e) adv v' (accounts st)) st''' = Normal (acc, st''')" by simp + define k\<^sub>o m\<^sub>o + where "k\<^sub>o = stack st'''" + and "m\<^sub>o = memory st'''" + show ?thesis + proof (cases "stmt fb e' emptyStore (st'''\accounts := acc, stack:=emptyStore, memory:=emptyStore\)") + case n2: (n a st'''') + with g_def a1 gcost n Pair KValue Value TAddr False Contract Some s2 p2 None n1 p3 k2 v s4 + have "stmt (EXTERNAL ad i xe val) e cd st = Normal ((), st''''\stack:=stack st''', memory := memory st'''\)" + using stmt.psimps(7)[OF 7(1)] st'_def st''_def st'''_def e'_def False s3 by simp + with a1 have "gas st6' \ gas st''''" by auto + also from 7(3)[OF l1 l2 l3 l4 l5 l6 l7 _ _ l8 l9 _ _ _ None l10, where ?s'k="st'''" and ?s'l="st'''\accounts := acc,stack:=emptyStore, memory:=emptyStore\" and ?x=e', of "(x0,fb)" x0 fb] n2 + have "\ \ gas (st'''\accounts := acc,stack:=emptyStore, memory:=emptyStore\)" using e'_def s3 by simp + also from msel_ssel_expr_load_rexp_gas(3)[of val e cd st'' "gas st''"] + have "\ \ gas st''" using n1 st'_def st''_def st'''_def p3 by simp + also from msel_ssel_expr_load_rexp_gas(3)[of ad e cd st' "gas st - g"] + have "\ \ gas st'" using n st'_def st''_def st'''_def p3 by fastforce + finally show ?thesis using st'_def by simp + next + case (e x) + with 7(1) g_def a1 gcost n Pair KValue Value TAddr False Some s2 Contract p2 None n1 p3 k2 v s4 s3 show ?thesis using stmt.psimps(7)[of ad i xe val e cd st] st'_def st''_def st'''_def e'_def by simp + qed + qed + next + case s1: (Some a) + then show ?thesis + proof (cases a) + case (Method x1) + then show ?thesis + proof (cases x1) + case p4: (fields fp ext f) + then show ?thesis + proof (cases ext) + case True + then show ?thesis + proof (cases "load True fp xe e' emptyStore emptyStore emptyStore e cd st''' (gas st''')") + case s4: (n a g''') + define st'''' where "st'''' = st'''\gas := g'''\" + then show ?thesis + proof (cases a) + case f1: (fields e\<^sub>l cd\<^sub>l k\<^sub>l m\<^sub>l) + then have l10: "toState (load True fp xe e' emptyStore emptyStore emptyStore e cd) st''' = Normal ((e\<^sub>l, cd\<^sub>l, k\<^sub>l, m\<^sub>l), st'''')" using s4 st''''_def by simp + show ?thesis + proof (cases "transfer (address e) adv v' (accounts st'''')") + case n2: None + with 7(1) g_def a1 gcost n Pair KValue Value TAddr False Some s2 Contract p2 s1 Method p4 n1 p3 k2 v s3 f1 e'_def True s4 show ?thesis using stmt.psimps(7)[of ad i xe val e cd st] st'_def st''_def st'''_def st''''_def by simp + next + case s5: (Some acc) + then have l11: "option Err (\st. transfer (address e) adv v' (accounts st)) st'''' = Normal (acc, st'''')" by simp + define k\<^sub>o where "k\<^sub>o = accounts st''''" + define m\<^sub>o where "m\<^sub>o = accounts st''''" + show ?thesis + proof (cases "stmt f e\<^sub>l cd\<^sub>l (st''''\accounts := acc, stack:=k\<^sub>l,memory:=m\<^sub>l\)") + case n2: (n a st''''') + with 7(1) g_def a1 gcost n Pair KValue Value TAddr Some s2 Contract p2 s1 Method p4 n1 p3 k2 v k\<^sub>o_def m\<^sub>o_def e'_def s3 f1 s4 s5 + have "stmt (EXTERNAL ad i xe val) e cd st = Normal ((), st'''''\stack:=stack st'''', memory := memory st''''\)" + using stmt.psimps(7)[of ad i xe val e cd st] st'_def st''_def st'''_def st''''_def True False by simp + with a1 have "gas st6' \ gas st'''''" by auto + also from 7(2)[OF l1 l2 l3 l4 l5 l6 l7 _ _ l8 l9 _ _ _ s1 Method _ _ _ l10 _ _ _ l11, where ?s'm="st''''\accounts := acc, stack := k\<^sub>l, memory := m\<^sub>l\" and ?s'l=st''''] p4 n2 e'_def True s3 + have "\ \ gas (st''''\accounts := acc, stack := k\<^sub>l, memory := m\<^sub>l\)" by simp + also from msel_ssel_expr_load_rexp_gas(4)[of True fp xe e' emptyStore emptyStore emptyStore e cd st''' "gas st'''"] + have "\ \ gas st'''" using s3 st'_def st''_def st'''_def st''''_def f1 s4 by simp + also from msel_ssel_expr_load_rexp_gas(3)[of val e cd st'' "gas st''"] + have "\ \ gas st''" using n1 st'_def st''_def st'''_def by fastforce + also from msel_ssel_expr_load_rexp_gas(3)[of ad e cd st' "gas st - g"] + have "\ \ gas st'" using n st'_def st''_def st'''_def by fastforce + finally show ?thesis using st'_def by simp + next + case (e x) + with 7(1) g_def a1 gcost n Pair KValue Value TAddr False Some s2 Contract p2 s1 Method p4 n1 p3 k2 v k\<^sub>o_def m\<^sub>o_def e'_def s3 f1 s4 s5 True show ?thesis using stmt.psimps(7) st'_def st''_def st'''_def st''''_def by simp + qed + qed + qed + next + case (e x) + with 7(1) g_def a1 gcost n Pair KValue Value TAddr False Some s2 Contract p2 s1 Method p4 n1 p3 k2 v e'_def True s3 show ?thesis using stmt.psimps(7) st'_def st''_def st'''_def by simp + qed + next + case f: False + with 7(1) g_def a1 gcost n Pair KValue Value TAddr False Some s2 Contract p2 s1 Method p4 n1 p3 k2 v s3 show ?thesis using stmt.psimps(7) st'_def st''_def st'''_def by simp + qed + qed + next + case (Function _) + with 7(1) g_def a1 gcost n Pair KValue Value TAddr False Some s2 Contract p2 s1 n1 p3 k2 v s3 show ?thesis using stmt.psimps(7) st'_def st''_def st'''_def by simp + next + case (Var _) + with 7(1) g_def a1 gcost n Pair KValue Value TAddr False Some s2 Contract p2 s1 n1 p3 k2 v s3 show ?thesis using stmt.psimps(7) st'_def st''_def st'''_def by simp + qed + qed + qed + next + case (Calldata x2) + with 7(1) g_def a1 gcost n Pair KValue Value TAddr False Some s2 Contract p2 n1 p3 k2 show ?thesis using stmt.psimps(7) st'_def st''_def st'''_def by simp + next + case (Memory x3) + with 7(1) g_def a1 gcost n Pair KValue Value TAddr False Some s2 Contract p2 n1 p3 k2 show ?thesis using stmt.psimps(7) st'_def st''_def st'''_def by simp + next + case (Storage x4) + with 7(1) g_def a1 gcost n Pair KValue Value TAddr False Some s2 Contract p2 n1 p3 k2 show ?thesis using stmt.psimps(7) st'_def st''_def st'''_def by simp + qed + next + case (KCDptr x2) + with 7(1) g_def a1 gcost n Pair KValue Value TAddr False Some s2 Contract p2 n1 p3 show ?thesis using stmt.psimps(7) st'_def st''_def st'''_def by simp + next + case (KMemptr x3) + with 7(1) g_def a1 gcost n Pair KValue Value TAddr False Some s2 Contract p2 n1 p3 show ?thesis using stmt.psimps(7) st'_def st''_def st'''_def by simp + next + case (KStoptr x4) + with 7(1) g_def a1 gcost n Pair KValue Value TAddr False Some s2 Contract p2 n1 p3 show ?thesis using stmt.psimps(7) st'_def st''_def st'''_def by simp + qed + qed + next + case n2: (e x) + with 7(1) g_def a1 gcost n Pair KValue Value TAddr False Some s2 Contract p2 show ?thesis using stmt.psimps(7) st'_def st''_def by simp + qed + qed + qed + qed + qed + qed + qed + next + case (Calldata x2) + with 7(1) g_def a1 gcost n Pair KValue show ?thesis using stmt.psimps(7) st'_def st''_def by simp + next + case (Memory x3) + with 7(1) g_def a1 gcost n Pair KValue show ?thesis using stmt.psimps(7) st'_def st''_def by simp + next + case (Storage x4) + with 7(1) g_def a1 gcost n Pair KValue show ?thesis using stmt.psimps(7) st'_def st''_def by simp + qed + next + case (KCDptr x2) + with 7(1) g_def a1 gcost n Pair show ?thesis using stmt.psimps(7) st'_def st''_def by simp + next + case (KMemptr x3) + with 7(1) g_def a1 gcost n Pair show ?thesis using stmt.psimps(7) st'_def st''_def by simp + next + case (KStoptr x4) + with 7(1) g_def a1 gcost n Pair show ?thesis using stmt.psimps(7) st'_def st''_def by simp + qed + qed + next + case (e _) + with 7(1) g_def a1 gcost show ?thesis using stmt.psimps(7) st'_def by simp + qed + qed + qed +next + case (8 ad ex e cd st) + define g where "g = costs (TRANSFER ad ex) e cd st" + show ?case + proof (rule allI[OF impI]) + fix st6' assume stmt_def: "stmt (TRANSFER ad ex) e cd st = Normal ((), st6')" + show "gas st6' \ gas st" + proof cases + assume "gas st \ g" + with 8 stmt_def g_def show ?thesis using stmt.psimps(8)[of ad ex e cd st] by simp + next + assume "\ gas st \ g" + then have l1: "assert Gas (\st. costs (TRANSFER ad ex) e cd st < gas st) st = Normal ((), st) " using g_def by simp + define st' where "st' = st\gas := gas st - g\" + then have l2: " modify (\st. st\gas := gas st - costs (TRANSFER ad ex) e cd st\) st = Normal ((), st')" using g_def by simp + show ?thesis + proof (cases "expr ad e cd st' (gas st - g)") + case (n a0 g') + define st'' where "st'' = st'\gas := g'\" + with n have l3: "toState (expr ad e cd) st' = Normal (a0, st'')" using st'_def by simp + then show ?thesis + proof (cases a0) + case (Pair b c) + then show ?thesis + proof (cases b) + case (KValue adv) + then show ?thesis + proof (cases c) + case (Value x1) + then show ?thesis + proof (cases x1) + case (TSInt x1) + with 8(1) stmt_def `\ gas st \ g` n Pair KValue Value g_def show ?thesis using stmt.psimps(8) st'_def st''_def by simp + next + case (TUInt x2) + with 8(1) stmt_def `\ gas st \ g` n Pair KValue Value g_def show ?thesis using stmt.psimps(8) st'_def st''_def by simp + next + case TBool + with 8(1) stmt_def `\ gas st \ g` n Pair KValue Value g_def show ?thesis using stmt.psimps(8) st'_def st''_def by simp + next + case TAddr + then have l4: "(case a0 of (KValue adv, Value TAddr) \ return adv | (KValue adv, Value _) \ throw Err | (KValue adv, _) \ throw Err + | (_, b) \ throw Err) st'' = Normal (adv, st'')" using Pair KValue Value by simp + then show ?thesis + proof (cases "expr ex e cd st'' (gas st'')") + case n2: (n a1 g'') + define st''' where "st''' = st''\gas := g''\" + with n2 have l5: "toState (expr ex e cd) st'' = Normal (a1, st''')" by simp + then show ?thesis + proof (cases a1) + case p2: (Pair b c) + then show ?thesis + proof (cases b) + case k2: (KValue v) + then show ?thesis + proof (cases c) + case v2: (Value t) + then have l6: "(case a1 of (KValue v, Value t) \ return (v, t) | (KValue v, _) \ throw Err | (_, b) \ throw Err) st''' = Normal ((v,t), st''')" using p2 k2 by simp + then show ?thesis + proof (cases "convert t (TUInt 256) v") + case None + with 8(1) stmt_def g_def `\ gas st \ g` n Pair KValue Value n2 p2 k2 v2 TAddr show ?thesis using stmt.psimps(8) st'_def st''_def st'''_def by simp + next + case (Some v') + then show ?thesis + proof (cases "type (accounts st''' adv)") + case None + with 8(1) stmt_def g_def `\ gas st \ g` n Pair KValue Value n2 p2 k2 v2 TAddr Some show ?thesis using stmt.psimps(8) st'_def st''_def st'''_def by simp + next + case s0: (Some a) + then show ?thesis + proof (cases a) + case EOA + then show ?thesis + proof (cases "transfer (address e) adv v' (accounts st''')") + case None + with 8(1) stmt_def g_def `\ gas st \ g` n Pair KValue Value n2 p2 k2 v2 TAddr Some EOA s0 show ?thesis using stmt.psimps(8) st'_def st''_def st'''_def by simp + next + case s1: (Some acc) + then have l7: "option Err (\st. transfer (address e) adv v' (accounts st)) st''' = Normal (acc, st''')" using st'''_def by simp + with 8(1) `\ gas st \ g` n Pair KValue Value n2 p2 k2 v2 TAddr Some EOA g_def s0 + have "stmt (TRANSFER ad ex) e cd st = Normal ((),st'''\accounts:=acc\)" using stmt.psimps(8)[of ad ex e cd st] st'_def st''_def st'''_def by simp + with stmt_def have "gas st6' = gas (st'''\accounts:=acc\)" by auto + also from msel_ssel_expr_load_rexp_gas(3)[of ex e cd st'' "gas st''"] + have "\ \ gas st''" using st'_def st''_def st'''_def n2 by fastforce + also from msel_ssel_expr_load_rexp_gas(3)[of ad e cd st' "gas st - g"] + have "\ \ gas st'" using st'_def st''_def n by fastforce + finally show ?thesis using st'_def by simp + qed + next + case (Contract c) + then show ?thesis + proof (cases "ep $$ c") + case None + with 8(1) stmt_def g_def `\ gas st \ g` n Pair KValue Value n2 p2 k2 v2 TAddr Contract Some s0 show ?thesis using stmt.psimps(8) st'_def st''_def st'''_def by simp + next + case s2: (Some a) + then show ?thesis + proof (cases a) + case p3: (fields ct cn f) + with s2 have l7: "option Err (\_. ep $$ c) st''' = Normal ((ct, cn, f), st''')" by simp + define e' where "e' = ffold_init ct (emptyEnv adv c (address e) v') (fmdom ct)" + show ?thesis + proof (cases "transfer (address e) adv v' (accounts st''')") + case None + with 8(1) stmt_def g_def `\ gas st \ g` n Pair KValue Value n2 p2 k2 v2 TAddr Contract Some s2 p3 s0 show ?thesis using stmt.psimps(8) st'_def st''_def st'''_def by simp + next + case s3: (Some acc) + then have l8: "option Err (\st. transfer (address e) adv v' (accounts st)) st''' = Normal (acc, st''')" by simp + then show ?thesis + proof (cases "stmt f e' emptyStore (st'''\accounts := acc, stack:=emptyStore, memory:=emptyStore\)") + case n3: (n a st'''') + with 8(1) `\ gas st \ g` n Pair KValue Value n2 p2 k2 v2 TAddr Some s2 p3 g_def Contract s3 s0 + have "stmt (TRANSFER ad ex) e cd st = Normal ((),st''''\stack:=stack st''', memory := memory st'''\)" using e'_def stmt.psimps(8)[of ad ex e cd st] st'_def st''_def st'''_def by simp + with stmt_def have "gas st6' \ gas st''''" by auto + also from 8(2)[OF l1 l2 l3 l4 l5 l6, of v t _ _ "accounts st'''" "st'''", OF _ _ _ s0 Contract l7 _ _ _ _ _ l8, where ?s'k="st'''\accounts := acc, stack := emptyStore, memory := emptyStore\"] `\ gas st \ g` e'_def n3 Some + have "\ \ gas (st'''\accounts := acc, stack := emptyStore, memory := emptyStore\)" by simp + also from msel_ssel_expr_load_rexp_gas(3)[of ex e cd st'' "gas st''"] + have "\ \ gas st''" using st'_def st''_def st'''_def n2 by fastforce + also from msel_ssel_expr_load_rexp_gas(3)[of ad e cd st' "gas st - g"] + have "\ \ gas st'" using st'_def st''_def n by fastforce + finally show ?thesis using st'_def by simp + next + case (e x) + with 8(1) `\ gas st \ g` n Pair KValue Value n2 p2 k2 v2 TAddr Some s2 p3 g_def e'_def stmt_def Contract s3 s0 show ?thesis using stmt.psimps(8)[of ad ex e cd st] st'_def st''_def st'''_def by simp + qed + qed + qed + qed + qed + qed + qed + next + case (Calldata x2) + with 8(1) stmt_def `\ gas st \ g` n Pair KValue Value TAddr n2 p2 k2 g_def show ?thesis using stmt.psimps(8) st'_def st''_def st'''_def by simp + next + case (Memory x3) + with 8(1) stmt_def `\ gas st \ g` n Pair KValue Value TAddr n2 p2 k2 g_def show ?thesis using stmt.psimps(8) st'_def st''_def st'''_def by simp + next + case (Storage x4) + with 8(1) stmt_def `\ gas st \ g` n Pair KValue Value TAddr n2 p2 k2 g_def show ?thesis using stmt.psimps(8) st'_def st''_def st'''_def by simp + qed + next + case (KCDptr x2) + with 8(1) stmt_def `\ gas st \ g` n Pair KValue Value TAddr n2 p2 g_def show ?thesis using stmt.psimps(8) st'_def st''_def st'''_def by simp + next + case (KMemptr x3) + with 8(1) stmt_def `\ gas st \ g` n Pair KValue Value TAddr n2 p2 g_def show ?thesis using stmt.psimps(8) st'_def st''_def st'''_def by simp + next + case (KStoptr x4) + with 8(1) stmt_def `\ gas st \ g` n Pair KValue Value TAddr n2 p2 g_def show ?thesis using stmt.psimps(8) st'_def st''_def st'''_def by simp + qed + qed + next + case (e e) + with 8(1) stmt_def `\ gas st \ g` n Pair KValue Value TAddr g_def show ?thesis using stmt.psimps(8) st'_def st''_def by simp + qed + qed + next + case (Calldata x2) + with 8(1) stmt_def `\ gas st \ g` n Pair KValue g_def show ?thesis using stmt.psimps(8) st'_def st''_def by simp + next + case (Memory x3) + with 8(1) stmt_def `\ gas st \ g` n Pair KValue g_def show ?thesis using stmt.psimps(8) st'_def st''_def by simp + next + case (Storage x4) + with 8(1) stmt_def `\ gas st \ g` n Pair KValue g_def show ?thesis using stmt.psimps(8) st'_def st''_def by simp + qed + next + case (KCDptr x2) + with 8(1) stmt_def `\ gas st \ g` n Pair g_def show ?thesis using stmt.psimps(8) st'_def st''_def by simp + next + case (KMemptr x3) + with 8(1) stmt_def `\ gas st \ g` n Pair g_def show ?thesis using stmt.psimps(8) st'_def st''_def by simp + next + case (KStoptr x4) + with 8(1) stmt_def `\ gas st \ g` n Pair g_def show ?thesis using stmt.psimps(8) st'_def st''_def by simp + qed + qed + next + case (e e) + with 8(1) stmt_def `\ gas st \ g` g_def show ?thesis using stmt.psimps(8) st'_def by simp + qed + qed + qed +next + case (9 id0 tp s e\<^sub>v cd st) + define g where "g = costs (BLOCK ((id0, tp), None) s) e\<^sub>v cd st" + show ?case + proof (rule allI[OF impI]) + fix st6' assume stmt_def: "stmt (BLOCK ((id0, tp), None) s) e\<^sub>v cd st = Normal ((), st6')" + show "gas st6' \ gas st" + proof cases + assume "gas st \ g" + with 9 stmt_def g_def show ?thesis using stmt.psimps(9) by simp + next + assume "\ gas st \ g" + then have l1: "assert Gas (\st. costs (BLOCK ((id0, tp), None) s) e\<^sub>v cd st < gas st) st = Normal ((), st) " using g_def by simp + define st' where "st' = st\gas := gas st - g\" + then have l2: "modify (\st. st\gas := gas st - costs (BLOCK ((id0, tp), None) s) e\<^sub>v cd st\) st = Normal ((), st')" using g_def by simp + show ?thesis + proof (cases "decl id0 tp None False cd (memory st') (storage st') (cd, (memory st'), (stack st'), e\<^sub>v)") + case n2: None + with 9 stmt_def `\ gas st \ g` g_def show ?thesis using stmt.psimps(9) st'_def by simp + next + case (Some a) + then have l3: "option Err (\st. decl id0 tp None False cd (memory st) (storage st) (cd, memory st, stack st, e\<^sub>v)) st' = Normal (a, st')" by simp + then show ?thesis + proof (cases a) + case (fields cd' mem' sck' e') + with 9(1) stmt_def `\ gas st \ g` g_def have "stmt (BLOCK ((id0, tp), None) s) e\<^sub>v cd st = stmt s e' cd' (st\gas := gas st - g, stack := sck', memory := mem'\)" using stmt.psimps(9)[OF 9(1)] Some st'_def by simp + with 9(2)[OF l1 l2 l3] stmt_def `\ gas st \ g` fields g_def have "gas st6' \ gas (st\gas := gas st - g, stack := sck', memory := mem'\)" using st'_def by fastforce + then show ?thesis by simp + qed + qed + qed + qed +next + case (10 id0 tp ex' s e\<^sub>v cd st) + define g where "g = costs (BLOCK ((id0, tp), Some ex') s) e\<^sub>v cd st" + show ?case + proof (rule allI[OF impI]) + fix st6' assume stmt_def: "stmt (BLOCK ((id0, tp), Some ex') s) e\<^sub>v cd st = Normal ((), st6')" + show "gas st6' \ gas st" + proof cases + assume "gas st \ g" + with 10 stmt_def g_def show ?thesis using stmt.psimps(10) by simp + next + assume "\ gas st \ g" + then have l1: "assert Gas (\st. costs (BLOCK ((id0, tp), Some ex') s) e\<^sub>v cd st < gas st) st = Normal ((), st) " using g_def by simp + define st' where "st' = st\gas := gas st - g\" + then have l2: "modify (\st. st\gas := gas st - costs (BLOCK ((id0, tp), Some ex') s) e\<^sub>v cd st\) st = Normal ((), st')" using g_def by simp + show ?thesis + proof (cases "expr ex' e\<^sub>v cd st' (gas st - g)") + case (n a g') + define st'' where "st'' = st'\gas := g'\" + with n have l3: "toState (expr ex' e\<^sub>v cd) st' = Normal (a, st'')" using st''_def st'_def by simp + then show ?thesis + proof (cases a) + case (Pair v t) + then show ?thesis + proof (cases "decl id0 tp (Some (v, t)) False cd (memory st'') (storage st'') (cd, memory st'', stack st'', e\<^sub>v)") + case None + with 10(1) stmt_def `\ gas st \ g` n Pair g_def show ?thesis using stmt.psimps(10) st'_def st''_def by simp + next + case s2: (Some a) + then have l4: "option Err (\st. decl id0 tp (Some (v, t)) False cd (memory st) (storage st) (cd, memory st, stack st, e\<^sub>v)) st'' = Normal (a, st'')" by simp + then show ?thesis + proof (cases a) + case (fields cd' mem' sck' e') + with 10(1) stmt_def `\ gas st \ g` n Pair s2 g_def have "stmt (BLOCK ((id0, tp), Some ex') s) e\<^sub>v cd st = stmt s e' cd' (st''\stack := sck', memory := mem'\)" using stmt.psimps(10)[of id0 tp ex' s e\<^sub>v cd st] st'_def st''_def by simp + with 10(2)[OF l1 l2 l3 Pair l4 fields, where s'd="st''\stack := sck', memory := mem'\"] n stmt_def `\ gas st \ g` s2 fields g_def have "gas st6' \ gas st''" by simp + moreover from msel_ssel_expr_load_rexp_gas(3)[of ex' e\<^sub>v cd st' "gas st - g"] n have "gas st'' \ gas st'" using st'_def st''_def by fastforce + ultimately show ?thesis using st'_def by simp + qed + qed + qed + next + case (e e) + with 10 stmt_def `\ gas st \ g` g_def show ?thesis using stmt.psimps(10) st'_def by simp + qed + qed + qed +next + case (11 i xe val e cd st) + define g where "g = costs (NEW i xe val) e cd st" + show ?case + proof (rule allI[OF impI]) + fix st6' assume a1: "stmt (NEW i xe val) e cd st = Normal ((), st6')" + show "gas st6' \ gas st" + proof (cases) + assume "gas st \ g" + with 11(1) a1 show ?thesis using stmt.psimps(11) g_def by simp + next + assume gcost: "\ gas st \ g" + then have l1: "assert Gas (\st. costs (NEW i xe val) e cd st < gas st) st = Normal ((), st) " using g_def by simp + define st' where "st' = st\gas := gas st - g\" + then have l2: "modify (\st. st\gas := gas st - costs (NEW i xe val) e cd st\) st = Normal ((), st')" using g_def by simp + define adv where "adv = hash (address e) (ShowL\<^sub>n\<^sub>a\<^sub>t (contracts (accounts st' (address e))))" + then show ?thesis + proof (cases "type (accounts st' adv) = None") + case True + then show ?thesis + proof (cases "expr val e cd st' (gas st')") + case n0: (n kv g') + define st'' where "st'' = st'\gas := g'\" + then have l4: "toState (expr val e cd) st' = Normal (kv, st'')" using n0 by simp + then show ?thesis + proof (cases kv) + case p0: (Pair a b) + then show ?thesis + proof (cases a) + case k0: (KValue v) + then show ?thesis + proof (cases b) + case v0: (Value t) + then show ?thesis + proof (cases "ep $$ i") + case None + with a1 gcost g_def True n0 p0 k0 v0 + show ?thesis using stmt.psimps(11)[OF 11(1)] adv_def st'_def st''_def by simp + next + case s0: (Some a) + then have l5: "option Err (\_. ep $$ i) st'' = Normal (a, st'')" by simp + then show ?thesis + proof (cases a) + case f0: (fields ct cn _) + define e' where "e' = ffold_init ct (emptyEnv adv i (address e) v) (fmdom ct)" + then show ?thesis + proof (cases "load True (fst cn) xe e' emptyStore emptyStore emptyStore e cd st'' (gas st'')") + case n1: (n a g'') + define st''' where "st''' = st''\gas := g''\" + then have l6: "toState (load True (fst cn) xe e' emptyStore emptyStore emptyStore e cd) st'' = Normal (a, st''')" using n1 by simp + then show ?thesis + proof (cases a) + case f1: (fields e\<^sub>l cd\<^sub>l k\<^sub>l m\<^sub>l) + define st'''' where "st'''' = st'''\accounts:=(accounts st''')(adv := \bal = ShowL\<^sub>i\<^sub>n\<^sub>t 0, type = Some (Contract i), contracts = 0\), storage:=(storage st''')(adv := {$$})\" + then show ?thesis + proof (cases "transfer (address e) adv v (accounts st'''')") + case None + with a1 gcost g_def True n0 p0 k0 v0 s0 f0 n1 f1 + show ?thesis using stmt.psimps(11)[OF 11(1)] adv_def e'_def st'_def st''_def st'''_def st''''_def by (simp add:Let_def) + next + case s1: (Some acc) + define st''''' where "st''''' = st''''\accounts := acc, stack:=k\<^sub>l, memory:=m\<^sub>l\" + then show ?thesis + proof (cases "stmt (snd cn) e\<^sub>l cd\<^sub>l st'''''") + case (n a st'''''') + define st''''''' where "st''''''' = st''''''\stack:=stack st'''', memory := memory st''''\" + define st'''''''' where "st'''''''' = incrementAccountContracts (address e) st'''''''" + from a1 gcost g_def True n0 p0 k0 v0 s0 f0 n1 f1 s1 n have "st6' = st''''''''" + using st'_def st''_def st'''_def st''''_def st'''''_def st'''''''_def st''''''''_def + stmt.psimps(11)[OF 11(1)] adv_def e'_def by (simp add:Let_def) + then have "gas st6' = gas st''''''''" by simp + also have "\ \ gas st'''''''" using st''''''''_def incrementAccountContracts_def by simp + also have "\ \ gas st''''''" using st'''''''_def by simp + also have "\ \ gas st'''''" using 11(2)[OF l1 l2 _ _ l4 _ _ l5 _ _ e'_def l6, where ?s'h="st''''" and ?s'i="st''''" and ?s'j="st''''" and ?s'k="st''''\accounts := acc, stack := k\<^sub>l, memory := m\<^sub>l\", of st' "()"] p0 k0 v0 f0 f1 s1 n True st''''_def st'''''_def adv_def by simp + also have "\ \ gas st''''" using st'''''_def by simp + also have "\ \ gas st'''" using st''''_def by simp + also have "\ \ gas st''" using st'''_def msel_ssel_expr_load_rexp_gas(4) n1 f1 by simp + also have "\ \ gas st'" using st''_def msel_ssel_expr_load_rexp_gas(3) n0 p0 by simp + also have "\ \ gas st" using st'_def by simp + finally show ?thesis . + next + case (e e) + with a1 gcost g_def n0 True p0 k0 v0 s0 f0 n1 f1 s1 + show ?thesis using stmt.psimps(11)[OF 11(1)] adv_def e'_def st'_def st''_def st'''_def st''''_def st'''''_def by (simp add:Let_def) + qed + qed + qed + next + case (e e) + with a1 gcost g_def n0 True p0 k0 v0 s0 f0 + show ?thesis using stmt.psimps(11)[OF 11(1)] adv_def e'_def st'_def st''_def by (simp add:Let_def) + qed + qed + qed + next + case (Calldata x2) + with a1 gcost g_def n0 True p0 k0 + show ?thesis using stmt.psimps(11)[OF 11(1)] adv_def st'_def by simp + next + case (Memory x3) + with a1 gcost g_def n0 True p0 k0 + show ?thesis using stmt.psimps(11)[OF 11(1)] adv_def st'_def by simp + next + case (Storage x4) + with a1 gcost g_def n0 True p0 k0 + show ?thesis using stmt.psimps(11)[OF 11(1)] adv_def st'_def by simp + qed + next + case (KCDptr x2) + with a1 gcost g_def n0 True p0 + show ?thesis using stmt.psimps(11)[OF 11(1)] adv_def st'_def by simp + next + case (KMemptr x3) + with a1 gcost g_def n0 True p0 + show ?thesis using stmt.psimps(11)[OF 11(1)] adv_def st'_def by simp + next + case (KStoptr x4) + with a1 gcost g_def n0 True p0 + show ?thesis using stmt.psimps(11)[OF 11(1)] adv_def st'_def by simp + qed + qed + next + case (e e) + with a1 gcost g_def True + show ?thesis using stmt.psimps(11)[OF 11(1)] adv_def st'_def by simp + qed + next + case False + with a1 gcost g_def + show ?thesis using stmt.psimps(11)[OF 11(1)] adv_def st'_def by (simp split:if_split_asm) + qed + qed + qed +qed + +subsection \Termination function\ + +text \Now we can prove termination using the lemma above.\ + +fun sgas + where "sgas l = gas (snd (snd (snd l)))" + +fun ssize + where "ssize l = size (fst l)" + +method stmt_dom_gas = + match premises in s: "stmt _ _ _ _ = Normal (_,_)" and d[thin]: "stmt_dom _" \ \insert stmt_dom_gas[OF d s]\ +method msel_ssel_expr_load_rexp = + match premises in e[thin]: "expr _ _ _ _ _ = Normal (_,_)" \ \insert msel_ssel_expr_load_rexp_gas(3)[OF e]\ | + match premises in l[thin]: "load _ _ _ _ _ _ _ _ _ _ _ = Normal (_,_)" \ \insert msel_ssel_expr_load_rexp_gas(4)[OF l, THEN conjunct1]\ +method costs = + match premises in "costs (WHILE ex s0) e cd st < _" for ex s0 and e::Environment and cd::CalldataT and st::State \ \insert while_not_zero[of (unchecked) ex s0 e cd st]\ | + match premises in "costs (INVOKE i xe) e cd st < _" for i xe and e::Environment and cd::CalldataT and st::State \ \insert invoke_not_zero[of (unchecked) i xe e cd st]\ | + match premises in "costs (EXTERNAL ad i xe val) e cd st < _" for ad i xe val and e::Environment and cd::CalldataT and st::State \ \insert external_not_zero[of (unchecked) ad i xe val e cd st]\ | + match premises in "costs (TRANSFER ad ex) e cd st < _" for ad ex and e::Environment and cd::CalldataT and st::State \ \insert transfer_not_zero[of (unchecked) ad ex e cd st]\ | + match premises in "costs (NEW i xe val) e cd st < _" for i xe val and e::Environment and cd::CalldataT and st::State \ \insert new_not_zero[of (unchecked) i xe val e cd st]\ + +termination stmt + apply (relation "measures [sgas, ssize]") + apply (auto split: if_split_asm result.split_asm Stackvalue.split_asm Type.split_asm Types.split_asm option.split_asm Member.split_asm bool.split_asm atype.split_asm) + apply ((stmt_dom_gas | msel_ssel_expr_load_rexp)+, costs?, simp)+ + done + +subsection \Gas Reduction\ + +text \ + The following corollary is a generalization of @{thm [source] msel_ssel_expr_load_rexp_dom_gas}. + We first prove that the function is defined for all input values and then obtain the final result as a corollary. +\ +lemma stmt_dom: "stmt_dom (s6, ev6, cd6, st6)" + apply (induct rule: stmt.induct[where ?P="\s6 ev6 cd6 st6. stmt_dom (s6, ev6, cd6, st6)"]) + apply (simp_all add: stmt.domintros(1-10)) + apply (rule stmt.domintros(11), force) + done + +lemmas stmt_gas = stmt_dom_gas[OF stmt_dom] + +lemma skip: + assumes "stmt SKIP ev cd st = Normal (x, st')" + shows "gas st > costs SKIP ev cd st" + and "st' = st\gas := gas st - costs SKIP ev cd st\" + using assms by (auto split:if_split_asm) + +lemma assign: + assumes "stmt (ASSIGN lv ex) ev cd st = Normal (xx, st')" + obtains (1) v t g l t' g' v' + where "expr ex ev cd (st\gas := gas st - costs (ASSIGN lv ex) ev cd st\) (gas st - costs (ASSIGN lv ex) ev cd st) = Normal ((KValue v, Value t), g)" + and "lexp lv ev cd (st\gas := g\) g = Normal((LStackloc l, Value t'),g')" + and "convert t t' v = Some v'" + and "st' = st\gas := g', stack := updateStore l (KValue v') (stack st)\" + | (2) v t g l t' g' v' + where "expr ex ev cd (st\gas := gas st - costs (ASSIGN lv ex) ev cd st\) (gas st - costs (ASSIGN lv ex) ev cd st) = Normal ((KValue v, Value t), g)" + and "lexp lv ev cd (st\gas := g\) g = Normal((LStoreloc l, Storage (STValue t')),g')" + and "convert t t' v = Some v'" + and "st' = st\gas := g', storage := (storage st) (address ev := (fmupd l v' (storage st (address ev))))\" + | (3) v t g l t' g' v' + where "expr ex ev cd (st\gas := gas st - costs (ASSIGN lv ex) ev cd st\) (gas st - costs (ASSIGN lv ex) ev cd st) = Normal ((KValue v, Value t), g)" + and "lexp lv ev cd (st\gas := g\) g = Normal((LMemloc l, Memory (MTValue t')),g')" + and "convert t t' v = Some v'" + and "st' = st\gas := g', memory := updateStore l (MValue v') (memory st)\" + | (4) p x t g l t' g' p' m + where "expr ex ev cd (st\gas := gas st - costs (ASSIGN lv ex) ev cd st\) (gas st - costs (ASSIGN lv ex) ev cd st) = Normal ((KCDptr p, Calldata (MTArray x t)), g)" + and "lexp lv ev cd (st\gas := g\) g = Normal((LStackloc l, Memory t'),g')" + and "accessStore l (stack st) = Some (KMemptr p')" + and "cpm2m p p' x t cd (memory st) = Some m" + and "st' = st\gas := g', memory := m\" + | (5) p x t g l t' g' p' s + where "expr ex ev cd (st\gas := gas st - costs (ASSIGN lv ex) ev cd st\) (gas st - costs (ASSIGN lv ex) ev cd st) = Normal ((KCDptr p, Calldata (MTArray x t)), g)" + and "lexp lv ev cd (st\gas := g\) g = Normal((LStackloc l, Storage t'),g')" + and "accessStore l (stack st) = Some (KStoptr p')" + and "cpm2s p p' x t cd (storage st (address ev)) = Some s" + and "st' = st\gas := g', storage := (storage st) (address ev := s)\" + | (6) p x t g l t' g' s + where "expr ex ev cd (st\gas := gas st - costs (ASSIGN lv ex) ev cd st\) (gas st - costs (ASSIGN lv ex) ev cd st) = Normal ((KCDptr p, Calldata (MTArray x t)), g)" + and "lexp lv ev cd (st\gas := g\) g = Normal((LStoreloc l, t'),g')" + and "cpm2s p l x t cd (storage st (address ev)) = Some s" + and "st' = st\gas := g', storage := (storage st) (address ev := s)\" + | (7) p x t g l t' g' m + where "expr ex ev cd (st\gas := gas st - costs (ASSIGN lv ex) ev cd st\) (gas st - costs (ASSIGN lv ex) ev cd st) = Normal ((KCDptr p, Calldata (MTArray x t)), g)" + and "lexp lv ev cd (st\gas := g\) g = Normal((LMemloc l, t'),g')" + and "cpm2m p l x t cd (memory st) = Some m" + and "st' = st\gas := g', memory := m\" + | (8) p x t g l t' g' + where "expr ex ev cd (st\gas := gas st - costs (ASSIGN lv ex) ev cd st\) (gas st - costs (ASSIGN lv ex) ev cd st) = Normal ((KMemptr p, Memory (MTArray x t)), g)" + and "lexp lv ev cd (st\gas := g\) g = Normal((LStackloc l, Memory t'),g')" + and "st' = st\gas := g', stack := updateStore l (KMemptr p) (stack st)\" + | (9) p x t g l t' g' p' s + where "expr ex ev cd (st\gas := gas st - costs (ASSIGN lv ex) ev cd st\) (gas st - costs (ASSIGN lv ex) ev cd st) = Normal ((KMemptr p, Memory (MTArray x t)), g)" + and "lexp lv ev cd (st\gas := g\) g = Normal((LStackloc l, Storage t'),g')" + and "accessStore l (stack st) = Some (KStoptr p')" + and "cpm2s p p' x t (memory st) (storage st (address ev)) = Some s" + and "st' = st\gas := g', storage := (storage st) (address ev := s)\" + | (10) p x t g l t' g' s + where "expr ex ev cd (st\gas := gas st - costs (ASSIGN lv ex) ev cd st\) (gas st - costs (ASSIGN lv ex) ev cd st) = Normal ((KMemptr p, Memory (MTArray x t)), g)" + and "lexp lv ev cd (st\gas := g\) g = Normal((LStoreloc l, t'),g')" + and "cpm2s p l x t (memory st) (storage st (address ev)) = Some s" + and "st' = st\gas := g', storage := (storage st) (address ev := s)\" + | (11) p x t g l t' g' + where "expr ex ev cd (st\gas := gas st - costs (ASSIGN lv ex) ev cd st\) (gas st - costs (ASSIGN lv ex) ev cd st) = Normal ((KMemptr p, Memory (MTArray x t)), g)" + and "lexp lv ev cd (st\gas := g\) g = Normal((LMemloc l, t'),g')" + and "st' = st\gas := g', memory := updateStore l (MPointer p) (memory st)\" + | (12) p x t g l t' g' p' m + where "expr ex ev cd (st\gas := gas st - costs (ASSIGN lv ex) ev cd st\) (gas st - costs (ASSIGN lv ex) ev cd st) = Normal ((KStoptr p, Storage (STArray x t)), g)" + and "lexp lv ev cd (st\gas := g\) g = Normal((LStackloc l, Memory t'),g')" + and "accessStore l (stack st) = Some (KMemptr p')" + and "cps2m p p' x t (storage st (address ev)) (memory st) = Some m" + and "st' = st\gas := g', memory := m\" + | (13) p x t g l t' g' + where "expr ex ev cd (st\gas := gas st - costs (ASSIGN lv ex) ev cd st\) (gas st - costs (ASSIGN lv ex) ev cd st) = Normal ((KStoptr p, Storage (STArray x t)), g)" + and "lexp lv ev cd (st\gas := g\) g = Normal((LStackloc l, Storage t'),g')" + and "st' = st\gas := g', stack := updateStore l (KStoptr p) (stack st)\" + | (14) p x t g l t' g' s + where "expr ex ev cd (st\gas := gas st - costs (ASSIGN lv ex) ev cd st\) (gas st - costs (ASSIGN lv ex) ev cd st) = Normal ((KStoptr p, Storage (STArray x t)), g)" + and "lexp lv ev cd (st\gas := g\) g = Normal((LStoreloc l, t'),g')" + and "copy p l x t (storage st (address ev)) = Some s" + and "st' = st\gas := g', storage := (storage st) (address ev := s)\" + | (15) p x t g l t' g' m + where "expr ex ev cd (st\gas := gas st - costs (ASSIGN lv ex) ev cd st\) (gas st - costs (ASSIGN lv ex) ev cd st) = Normal ((KStoptr p, Storage (STArray x t)), g)" + and "lexp lv ev cd (st\gas := g\) g = Normal((LMemloc l, t'),g')" + and "cps2m p l x t (storage st (address ev)) (memory st) = Some m" + and "st' = st\gas := g', memory := m\" + | (16) p t t' g l t'' g' + where "expr ex ev cd (st\gas := gas st - costs (ASSIGN lv ex) ev cd st\) (gas st - costs (ASSIGN lv ex) ev cd st) = Normal ((KStoptr p, Storage (STMap t t')), g)" + and "lexp lv ev cd (st\gas := g\) g = Normal((LStackloc l, t''),g')" + and "st' = st\gas := g', stack := updateStore l (KStoptr p) (stack st)\" +proof - + from assms consider + (1) v t g where "expr ex ev cd (st\gas := gas st - costs (ASSIGN lv ex) ev cd st\) (gas st - costs (ASSIGN lv ex) ev cd st) = Normal ((KValue v, Value t), g)" + | (2) p x t g where "expr ex ev cd (st\gas := gas st - costs (ASSIGN lv ex) ev cd st\) (gas st - costs (ASSIGN lv ex) ev cd st) = Normal ((KCDptr p, Calldata (MTArray x t)), g)" + | (3) p x t g where "expr ex ev cd (st\gas := gas st - costs (ASSIGN lv ex) ev cd st\) (gas st - costs (ASSIGN lv ex) ev cd st) = Normal ((KMemptr p, Memory (MTArray x t)), g)" + | (4) p x t g where "expr ex ev cd (st\gas := gas st - costs (ASSIGN lv ex) ev cd st\) (gas st - costs (ASSIGN lv ex) ev cd st) = Normal ((KStoptr p, Storage (STArray x t)), g)" + | (5) p t t' g where "expr ex ev cd (st\gas := gas st - costs (ASSIGN lv ex) ev cd st\) (gas st - costs (ASSIGN lv ex) ev cd st) = Normal ((KStoptr p, Storage (STMap t t')), g)" + by (auto split:if_split_asm result.split_asm Stackvalue.split_asm Type.split_asm MTypes.split_asm STypes.split_asm) + then show ?thesis + proof cases + case 1 + with assms consider + (11) l t' g' where "lexp lv ev cd (st\gas := g\) g = Normal((LStackloc l, Value t'),g')" + | (12) l t' g' where "lexp lv ev cd (st\gas := g\) g = Normal((LStoreloc l, Storage (STValue t')),g')" + | (13) l t' g' where "lexp lv ev cd (st\gas := g\) g = Normal((LMemloc l, Memory (MTValue t')),g')" + by (auto split:if_split_asm result.split_asm Type.split_asm LType.split_asm MTypes.split_asm STypes.split_asm) + then show ?thesis + proof cases + case 11 + with 1 assms show ?thesis using that(1) by (auto split:if_split_asm result.split_asm Type.split_asm LType.split_asm MTypes.split_asm STypes.split_asm option.split_asm) + next + case 12 + with 1 assms show ?thesis using that(2) by (auto split:if_split_asm result.split_asm Type.split_asm LType.split_asm MTypes.split_asm STypes.split_asm option.split_asm) + next + case 13 + with 1 assms show ?thesis using that(3) by (auto split:if_split_asm result.split_asm Type.split_asm LType.split_asm MTypes.split_asm STypes.split_asm option.split_asm) + qed + next + case 2 + with assms consider + (21) l t' g' where "lexp lv ev cd (st\gas := g\) g = Normal((LStackloc l, Memory t'),g')" + | (22) l t' g' where "lexp lv ev cd (st\gas := g\) g = Normal((LStackloc l, Storage t'),g')" + | (23) l t' g' where "lexp lv ev cd (st\gas := g\) g = Normal((LStoreloc l, t'),g')" + | (24) l t' g' where "lexp lv ev cd (st\gas := g\) g = Normal((LMemloc l, t'),g')" + by (auto split:if_split_asm result.split_asm Type.split_asm LType.split_asm MTypes.split_asm STypes.split_asm) + then show ?thesis + proof cases + case 21 + moreover from assms 2 21 obtain p' where 3: "accessStore l (stack st) = Some (KMemptr p')" + by (auto split:if_split_asm result.split_asm Type.split_asm LType.split_asm MTypes.split_asm STypes.split_asm option.split_asm Stackvalue.split_asm) + moreover from assms 2 21 3 obtain m where "cpm2m p p' x t cd (memory st) = Some m" + by (auto split:if_split_asm result.split_asm Type.split_asm LType.split_asm MTypes.split_asm STypes.split_asm option.split_asm Stackvalue.split_asm) + ultimately show ?thesis using that(4) assms 2 21 + by (auto split:if_split_asm result.split_asm Type.split_asm LType.split_asm MTypes.split_asm STypes.split_asm option.split_asm Stackvalue.split_asm) + next + case 22 + moreover from assms 2 22 obtain p' where 3: "accessStore l (stack st) = Some (KStoptr p')" + by (auto split:if_split_asm result.split_asm Type.split_asm LType.split_asm MTypes.split_asm STypes.split_asm option.split_asm Stackvalue.split_asm) + moreover from assms 2 22 3 4 obtain s where "cpm2s p p' x t cd (storage st (address ev)) = Some s" + by (auto split:if_split_asm result.split_asm Type.split_asm LType.split_asm MTypes.split_asm STypes.split_asm option.split_asm Stackvalue.split_asm) + ultimately show ?thesis using that(5) assms 2 22 + by (auto split:if_split_asm result.split_asm Type.split_asm LType.split_asm MTypes.split_asm STypes.split_asm option.split_asm Stackvalue.split_asm) + next + case 23 + moreover from assms 2 23 3 4 obtain s where "cpm2s p l x t cd (storage st (address ev)) = Some s" + by (auto split:if_split_asm result.split_asm Type.split_asm LType.split_asm MTypes.split_asm STypes.split_asm option.split_asm Stackvalue.split_asm) + ultimately show ?thesis using that(6) assms 2 + by (auto split:if_split_asm result.split_asm Type.split_asm LType.split_asm MTypes.split_asm STypes.split_asm option.split_asm Stackvalue.split_asm) + next + case 24 + moreover from assms 2 24 obtain m where "cpm2m p l x t cd (memory st) = Some m" + by (auto split:if_split_asm result.split_asm Type.split_asm LType.split_asm MTypes.split_asm STypes.split_asm option.split_asm Stackvalue.split_asm) + ultimately show ?thesis using that(7) assms 2 + by (auto split:if_split_asm result.split_asm Type.split_asm LType.split_asm MTypes.split_asm STypes.split_asm option.split_asm Stackvalue.split_asm) + qed + next + case 3 + with assms consider + (31) l t' g' where "lexp lv ev cd (st\gas := g\) g = Normal((LStackloc l, Memory t'),g')" + | (32) l t' g' where "lexp lv ev cd (st\gas := g\) g = Normal((LStackloc l, Storage t'),g')" + | (33) l t' g' where "lexp lv ev cd (st\gas := g\) g = Normal((LStoreloc l, t'),g')" + | (34) l t' g' where "lexp lv ev cd (st\gas := g\) g = Normal((LMemloc l, t'),g')" + by (auto split:if_split_asm result.split_asm Type.split_asm LType.split_asm MTypes.split_asm STypes.split_asm) + then show ?thesis + proof cases + case 31 + then show ?thesis using that(8) assms 3 by (auto split:if_split_asm) + next + case 32 + moreover from assms 3 32 obtain p' where 4: "accessStore l (stack st) = Some (KStoptr p')" + by (auto split:if_split_asm result.split_asm Type.split_asm LType.split_asm MTypes.split_asm STypes.split_asm option.split_asm Stackvalue.split_asm) + moreover from assms 3 32 4 5 obtain s where "cpm2s p p' x t (memory st) (storage st (address ev)) = Some s" + by (auto split:if_split_asm result.split_asm Type.split_asm LType.split_asm MTypes.split_asm STypes.split_asm option.split_asm Stackvalue.split_asm) + ultimately show ?thesis using that(9) assms 3 + by (auto split:if_split_asm result.split_asm Type.split_asm LType.split_asm MTypes.split_asm STypes.split_asm option.split_asm Stackvalue.split_asm) + next + case 33 + moreover from assms 3 33 3 4 obtain s where "cpm2s p l x t (memory st) (storage st (address ev)) = Some s" + by (auto split:if_split_asm result.split_asm Type.split_asm LType.split_asm MTypes.split_asm STypes.split_asm option.split_asm Stackvalue.split_asm) + ultimately show ?thesis using that(10) assms 3 + by (auto split:if_split_asm result.split_asm Type.split_asm LType.split_asm MTypes.split_asm STypes.split_asm option.split_asm Stackvalue.split_asm) + next + case 34 + then show ?thesis using that(11) assms 3 + by (auto split:if_split_asm result.split_asm Type.split_asm LType.split_asm MTypes.split_asm STypes.split_asm option.split_asm Stackvalue.split_asm) + qed + next + case 4 + with assms consider + (41) l t' g' where "lexp lv ev cd (st\gas := g\) g = Normal((LStackloc l, Memory t'),g')" + | (42) l t' g' where "lexp lv ev cd (st\gas := g\) g = Normal((LStackloc l, Storage t'),g')" + | (43) l t' g' where "lexp lv ev cd (st\gas := g\) g = Normal((LStoreloc l, t'),g')" + | (44) l t' g' where "lexp lv ev cd (st\gas := g\) g = Normal((LMemloc l, t'),g')" + by (auto split:if_split_asm result.split_asm Type.split_asm LType.split_asm MTypes.split_asm STypes.split_asm) + then show ?thesis + proof cases + case 41 + moreover from assms 4 41 obtain p' where 5: "accessStore l (stack st) = Some (KMemptr p')" + by (auto split:if_split_asm result.split_asm Type.split_asm LType.split_asm MTypes.split_asm STypes.split_asm option.split_asm Stackvalue.split_asm) + moreover from assms 4 41 5 6 obtain m where "cps2m p p' x t (storage st (address ev)) (memory st) = Some m" + by (auto split:if_split_asm result.split_asm Type.split_asm LType.split_asm MTypes.split_asm STypes.split_asm option.split_asm Stackvalue.split_asm) + ultimately show ?thesis using that(12) assms 4 + by (auto split:if_split_asm result.split_asm Type.split_asm LType.split_asm MTypes.split_asm STypes.split_asm option.split_asm Stackvalue.split_asm) + next + case 42 + then show ?thesis using that(13) assms 4 + by (auto split:if_split_asm result.split_asm Type.split_asm LType.split_asm MTypes.split_asm STypes.split_asm option.split_asm Stackvalue.split_asm) + next + case 43 + moreover from assms 4 43 5 obtain s where "copy p l x t (storage st (address ev)) = Some s" + by (auto split:if_split_asm result.split_asm Type.split_asm LType.split_asm MTypes.split_asm STypes.split_asm option.split_asm Stackvalue.split_asm) + ultimately show ?thesis using that(14) assms 4 + by (auto split:if_split_asm result.split_asm Type.split_asm LType.split_asm MTypes.split_asm STypes.split_asm option.split_asm Stackvalue.split_asm) + next + case 44 + moreover from assms 4 44 5 obtain m where "cps2m p l x t (storage st (address ev)) (memory st) = Some m" + by (auto split:if_split_asm result.split_asm Type.split_asm LType.split_asm MTypes.split_asm STypes.split_asm option.split_asm Stackvalue.split_asm) + ultimately show ?thesis using that(15) assms 4 + by (auto split:if_split_asm result.split_asm Type.split_asm LType.split_asm MTypes.split_asm STypes.split_asm option.split_asm Stackvalue.split_asm) + qed + next + case 5 + then show ?thesis using that(16) assms + by (auto split:if_split_asm result.split_asm Type.split_asm LType.split_asm MTypes.split_asm STypes.split_asm option.split_asm Stackvalue.split_asm) + qed +qed + +lemma comp: + assumes "stmt (COMP s1 s2) ev cd st = Normal (x, st')" + obtains (1) st'' + where "gas st > costs (COMP s1 s2) ev cd st" + and "stmt s1 ev cd (st\gas := gas st - costs (COMP s1 s2) ev cd st\) = Normal((), st'')" + and "stmt s2 ev cd st'' = Normal((), st')" + using assms by (simp split:if_split_asm result.split_asm prod.split_asm) + +lemma ite: + assumes "stmt (ITE ex s1 s2) ev cd st = Normal (x, st')" + obtains (True) g + where "gas st > costs (ITE ex s1 s2) ev cd st" + and "expr ex ev cd (st\gas := gas st - costs (ITE ex s1 s2) ev cd st\) (gas st - costs (ITE ex s1 s2) ev cd st) = Normal((KValue (ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True), Value TBool), g)" + and "stmt s1 ev cd (st\gas := g\) = Normal((), st')" +| (False) g + where "gas st > costs (ITE ex s1 s2) ev cd st" + and "expr ex ev cd (st\gas := gas st - costs (ITE ex s1 s2) ev cd st\) (gas st - costs (ITE ex s1 s2) ev cd st) = Normal((KValue (ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l False), Value TBool), g)" + and "stmt s2 ev cd (st\gas := g\) = Normal((), st')" + using assms by (simp split:if_split_asm result.split_asm prod.split_asm Stackvalue.split_asm Type.split_asm Types.split_asm) + +lemma while: + assumes "stmt (WHILE ex s0) ev cd st = Normal (x, st')" + obtains (True) g st'' + where "gas st > costs (WHILE ex s0) ev cd st" + and "expr ex ev cd (st\gas := gas st - costs (WHILE ex s0) ev cd st\) (gas st - costs (WHILE ex s0) ev cd st) = Normal((KValue (ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True), Value TBool), g)" + and "stmt s0 ev cd (st\gas := g\) = Normal((), st'')" + and "stmt (WHILE ex s0) ev cd st'' = Normal ((), st')" + | (False) g + where "gas st > costs (WHILE ex s0) ev cd st" + and "expr ex ev cd (st\gas := gas st - costs (WHILE ex s0) ev cd st\) (gas st - costs (WHILE ex s0) ev cd st) = Normal((KValue (ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l False), Value TBool), g)" + and "st' = st\gas := g\" + using assms +proof - + from assms have 1: "gas st > costs (WHILE ex s0) ev cd st" by (simp split:if_split_asm) + moreover from assms 1 have 2: "modify (\st. st\gas := gas st - costs (WHILE ex s0) ev cd st\) st = Normal ((), st\gas := gas st - costs (WHILE ex s0) ev cd st\)" by simp + moreover from assms 1 2 obtain b g where 3: "expr ex ev cd (st\gas := gas st - costs (WHILE ex s0) ev cd st\) (gas st - costs (WHILE ex s0) ev cd st) = Normal ((KValue b, Value TBool), g)" by (simp split:result.split_asm prod.split_asm Stackvalue.split_asm Type.split_asm Types.split_asm) + ultimately consider (True) "b = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True" | (False) "b = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l False" | (None) "b \ ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True \ b \ ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l False" by auto + then show ?thesis + proof cases + case True + moreover from assms 1 2 3 True obtain st' where 4: "stmt s0 ev cd (st\gas := g\) = Normal ((), st')" by (simp split:result.split_asm prod.split_asm) + moreover from assms 1 2 3 4 True obtain st'' where 5: "stmt (WHILE ex s0) ev cd st' = Normal ((), st'')" by (simp split:result.split_asm prod.split_asm) + ultimately show ?thesis using 1 2 3 that(1) assms by simp + next + case False + then show ?thesis using 1 2 3 that(2) assms true_neq_false by simp + next + case None + then show ?thesis using 1 2 3 assms by simp + qed +qed + +lemma invoke: + fixes ev + defines "e' members \ ffold (init members) (emptyEnv (address ev) (contract ev) (sender ev) (svalue ev)) (fmdom members)" + assumes "stmt (INVOKE i xe) ev cd st = Normal (x, st')" + obtains ct fb fp f e\<^sub>l cd\<^sub>l k\<^sub>l m\<^sub>l g st'' + where "gas st > costs (INVOKE i xe) ev cd st" + and "ep $$ contract ev = Some (ct, fb)" + and "ct $$ i = Some (Method (fp, False, f))" + and "load False fp xe (e' ct) emptyStore emptyStore (memory (st\gas := gas st - costs (INVOKE i xe) ev cd st\)) ev cd (st\gas := gas st - costs (INVOKE i xe) ev cd st\) (gas st - costs (INVOKE i xe) ev cd st) = Normal ((e\<^sub>l, cd\<^sub>l, k\<^sub>l, m\<^sub>l), g)" + and "stmt f e\<^sub>l cd\<^sub>l (st\gas:= g, stack:=k\<^sub>l, memory:=m\<^sub>l\) = Normal ((), st'')" + and "st' = st''\stack:=stack st\" +proof - + from assms have 1: "gas st > costs (INVOKE i xe) ev cd st" by (simp split:if_split_asm) + moreover from assms 1 obtain ct fb where 2: "ep $$ (contract ev) = Some (ct, fb)" by (simp split: prod.split_asm result.split_asm option.split_asm) + moreover from assms 1 2 obtain fp f where 3: "ct $$ i = Some (Method (fp, False, f))" by (simp split: prod.split_asm result.split_asm option.split_asm Member.split_asm bool.split_asm) + moreover from assms 1 2 3 obtain e\<^sub>l cd\<^sub>l k\<^sub>l m\<^sub>l g where 4: "load False fp xe (e' ct) emptyStore emptyStore (memory (st\gas := gas st - costs (INVOKE i xe) ev cd st\)) ev cd (st\gas := gas st - costs (INVOKE i xe) ev cd st\) (gas st - costs (INVOKE i xe) ev cd st) = Normal ((e\<^sub>l, cd\<^sub>l, k\<^sub>l, m\<^sub>l), g)" by (simp split: prod.split_asm result.split_asm) + moreover from assms 1 2 3 4 obtain st'' where 5: "stmt f e\<^sub>l cd\<^sub>l (st\gas:= g, stack:=k\<^sub>l, memory:=m\<^sub>l\) = Normal ((), st'')" by (simp split: prod.split_asm result.split_asm) + moreover from assms 1 2 3 4 5 have "st' = st''\stack:=stack st\" by (simp split: prod.split_asm result.split_asm) + ultimately show ?thesis using that by simp +qed + +lemma external: + fixes ev + defines "e' members adv c v \ ffold (init members) (emptyEnv adv c (address ev) v) (fmdom members)" + assumes "stmt (EXTERNAL ad' i xe val) ev cd st = Normal (x, st')" + obtains (Some) adv c g ct cn fb' v t g' v' fp f e\<^sub>l cd\<^sub>l k\<^sub>l m\<^sub>l g'' acc st'' + where "gas st > costs (EXTERNAL ad' i xe val) ev cd st" + and "expr ad' ev cd (st\gas := gas st - costs (EXTERNAL ad' i xe val) ev cd st\) (gas st - costs (EXTERNAL ad' i xe val) ev cd st) = Normal ((KValue adv, Value TAddr), g)" + and "adv \ address ev" + and "type (accounts (st\gas := g\) adv) = Some (Contract c)" + and "ep $$ c = Some (ct, cn, fb')" + and "expr val ev cd (st\gas := g\) g = Normal ((KValue v, Value t), g')" + and "convert t (TUInt 256) v = Some v'" + and "fmlookup ct i = Some (Method (fp, True, f))" + and "load True fp xe (e' ct adv c v') emptyStore emptyStore emptyStore ev cd (st\gas := g'\) g' = Normal ((e\<^sub>l, cd\<^sub>l, k\<^sub>l, m\<^sub>l), g'')" + and "transfer (address ev) adv v' (accounts (st\gas := g''\)) = Some acc" + and "stmt f e\<^sub>l cd\<^sub>l (st\gas := g'', accounts := acc, stack:=k\<^sub>l, memory:=m\<^sub>l\) = Normal ((), st'')" + and "st' = st''\stack:=stack st, memory := memory st\" + | (None) adv c g ct cn fb' v t g' v' acc st'' + where "gas st > costs (EXTERNAL ad' i xe val) ev cd st" + and "expr ad' ev cd (st\gas := gas st - costs (EXTERNAL ad' i xe val) ev cd st\) (gas st - costs (EXTERNAL ad' i xe val) ev cd st) = Normal ((KValue adv, Value TAddr), g)" + and "adv \ address ev" + and "type (accounts (st\gas := g\) adv) = Some (Contract c)" + and "ep $$ c = Some (ct, cn, fb')" + and "expr val ev cd (st\gas := g\) g = Normal ((KValue v, Value t), g')" + and "convert t (TUInt 256) v = Some v'" + and "ct $$ i = None" + and "transfer (address ev) adv v' (accounts st) = Some acc" + and "stmt fb' (e' ct adv c v') emptyStore (st\gas := g', accounts := acc, stack:=emptyStore, memory:=emptyStore\) = Normal ((), st'')" + and "st' = st''\stack:=stack st, memory := memory st\" +proof - + from assms have 1: "gas st > costs (EXTERNAL ad' i xe val) ev cd st" by (simp split:if_split_asm) + moreover from assms 1 obtain adv g where 2: "expr ad' ev cd (st\gas := gas st - costs (EXTERNAL ad' i xe val) ev cd st\) (gas st - costs (EXTERNAL ad' i xe val) ev cd st) = Normal ((KValue adv, Value TAddr), g)" by (simp split: prod.split_asm result.split_asm Stackvalue.split_asm Type.split_asm Types.split_asm) + moreover from assms 1 2 obtain c where 3: "type (accounts (st\gas := g\) adv) = Some (Contract c)" by (simp split: if_split_asm prod.split_asm result.split_asm Stackvalue.split_asm Type.split_asm Types.split_asm option.split_asm atype.split_asm) + moreover from assms 1 2 3 obtain ct cn fb' where 4: "ep $$ c = Some (ct, cn, fb')" by (simp add: Let_def split: if_split_asm prod.split_asm result.split_asm Stackvalue.split_asm Type.split_asm Types.split_asm option.split_asm) + moreover from assms 1 2 3 4 obtain v t g' where 5: "expr val ev cd (st\gas := g\) g = Normal ((KValue v, Value t), g')" using 1 2 by (simp split: if_split_asm prod.split_asm result.split_asm Stackvalue.split_asm Type.split_asm Types.split_asm option.split_asm) + moreover from assms 1 2 3 4 5 have 6: "adv \ address ev" by (simp add: Let_def split: if_split_asm prod.split_asm result.split_asm Stackvalue.split_asm Type.split_asm Types.split_asm option.split_asm) + moreover from assms 1 2 3 4 5 6 obtain v' where 7: "convert t (TUInt 256) v = Some v'" by (simp add: Let_def split: if_split_asm prod.split_asm result.split_asm Stackvalue.split_asm Type.split_asm Types.split_asm option.split_asm) + ultimately consider (Some) fp f where "ct $$ i = Some (Method (fp, True, f))" | (None) "fmlookup ct i = None" using assms by (simp add: Let_def split: if_split_asm prod.split_asm result.split_asm Stackvalue.split_asm Type.split_asm Types.split_asm option.split_asm Member.split_asm bool.split_asm) + then show ?thesis + proof cases + case (Some fp f) + moreover from assms 1 2 3 4 5 6 7 Some obtain e\<^sub>l cd\<^sub>l k\<^sub>l m\<^sub>l g'' where 8: "load True fp xe (e' ct adv c v') emptyStore emptyStore emptyStore ev cd (st\gas := g'\) g' = Normal ((e\<^sub>l, cd\<^sub>l, k\<^sub>l, m\<^sub>l), g'')" by (simp add: Let_def split: if_split_asm prod.split_asm result.split_asm Stackvalue.split_asm Type.split_asm Types.split_asm option.split_asm Member.split_asm) + moreover from assms 1 2 3 4 5 6 7 Some 8 obtain acc where 9: "transfer (address ev) adv v' (accounts st) = Some acc" by (simp add: Let_def split: if_split_asm prod.split_asm result.split_asm Stackvalue.split_asm Type.split_asm Types.split_asm option.split_asm Member.split_asm) + moreover from assms 1 2 3 4 5 6 7 Some 8 9 obtain st'' where 10: "stmt f e\<^sub>l cd\<^sub>l (st\gas := g'', accounts := acc, stack:=k\<^sub>l, memory:=m\<^sub>l\) = Normal ((), st'')" by (simp add: Let_def transfer_def split: if_split_asm prod.split_asm result.split_asm Stackvalue.split_asm Type.split_asm Types.split_asm option.split_asm Member.split_asm) + moreover from assms 1 2 3 4 5 6 7 Some 8 9 10 have "st' = st''\stack:=stack st, memory := memory st\" by (simp add: Let_def transfer_def split: if_split_asm prod.split_asm result.split_asm Stackvalue.split_asm Type.split_asm Types.split_asm option.split_asm Member.split_asm) + ultimately show ?thesis using 1 2 3 4 5 6 7 that(1) by simp + next + case None + moreover from assms 1 2 3 4 5 6 7 None obtain acc where 8: "transfer (address ev) adv v' (accounts st) = Some acc" by (simp add: Let_def split: if_split_asm prod.split_asm result.split_asm Stackvalue.split_asm Type.split_asm Types.split_asm option.split_asm Member.split_asm) + moreover from assms 1 2 3 4 5 6 7 None 8 obtain st'' where 9: "stmt fb' (e' ct adv c v') emptyStore (st\gas := g', accounts := acc, stack:=emptyStore, memory:=emptyStore\) = Normal ((), st'')" by (simp add: Let_def transfer_def split: if_split_asm prod.split_asm result.split_asm Stackvalue.split_asm Type.split_asm Types.split_asm option.split_asm Member.split_asm) + moreover from assms 1 2 3 4 5 6 7 None 8 9 have "st' = st''\stack:=stack st, memory := memory st\" by (simp add: Let_def transfer_def split: if_split_asm prod.split_asm result.split_asm Stackvalue.split_asm Type.split_asm Types.split_asm option.split_asm Member.split_asm) + ultimately show ?thesis using 1 2 3 4 5 6 7 that(2) by simp + qed +qed + +lemma transfer: + fixes ev + defines "e' members adv c st v \ ffold (init members) (emptyEnv adv c (address ev) v) (fmdom members)" + assumes "stmt (TRANSFER ad ex) ev cd st = Normal (x, st')" + obtains (Contract) v t g adv c g' v' acc ct cn f st'' + where "gas st > costs (TRANSFER ad ex) ev cd st" + and "expr ad ev cd (st\gas := gas st - costs (TRANSFER ad ex) ev cd st\) (gas st - costs (TRANSFER ad ex) ev cd st) = Normal ((KValue adv, Value TAddr), g)" + and "expr ex ev cd (st\gas := g\) g = Normal ((KValue v, Value t), g')" + and "convert t (TUInt 256) v = Some v'" + and "type (accounts (st\gas := g\) adv) = Some (Contract c)" + and "ep $$ c = Some (ct, cn, f)" + and "transfer (address ev) adv v' (accounts st) = Some acc" + and "stmt f (e' ct adv c (st\gas := g'\) v') emptyStore (st\gas := g', accounts := acc, stack:=emptyStore, memory:=emptyStore\) = Normal ((), st'')" + and "st' = st''\stack:=stack st, memory := memory st\" + | (EOA) v t g adv g' v' acc + where "gas st > costs (TRANSFER ad ex) ev cd st" + and "expr ad ev cd (st\gas := gas st - costs (TRANSFER ad ex) ev cd st\) (gas st - costs (TRANSFER ad ex) ev cd st) = Normal ((KValue adv, Value TAddr), g)" + and "expr ex ev cd (st\gas := g\) g = Normal ((KValue v, Value t), g')" + and "convert t (TUInt 256) v = Some v'" + and "type (accounts (st\gas := g\) adv) = Some EOA" + and "transfer (address ev) adv v' (accounts st) = Some acc" + and "st' = st\gas:=g', accounts:=acc\" +proof - + from assms have 1: "gas st > costs (TRANSFER ad ex) ev cd st" by (simp split:if_split_asm) + moreover from assms 1 obtain adv g where 2: "expr ad ev cd (st\gas := gas st - costs (TRANSFER ad ex) ev cd st\) (gas st - costs (TRANSFER ad ex) ev cd st) = Normal ((KValue adv, Value TAddr), g)" by (simp add: Let_def split: if_split_asm prod.split_asm result.split_asm Stackvalue.split_asm Type.split_asm Types.split_asm) + moreover from assms 1 2 obtain v t g' where 3: "expr ex ev cd (st\gas := g\) g = Normal ((KValue v, Value t), g')" by (simp add: Let_def split: if_split_asm prod.split_asm result.split_asm Stackvalue.split_asm Type.split_asm Types.split_asm) + moreover from assms 1 2 3 obtain v' where 4: "convert t (TUInt 256) v = Some v'" by (simp add: Let_def split: if_split_asm prod.split_asm result.split_asm Stackvalue.split_asm Type.split_asm Types.split_asm option.split_asm) + ultimately consider (Contract) c where "type (accounts (st\gas := g'\) adv) = Some (Contract c)" | (EOA) "type (accounts (st\gas := g'\) adv) = Some EOA" using assms by (simp add: Let_def split: if_split_asm prod.split_asm result.split_asm Stackvalue.split_asm Type.split_asm Types.split_asm option.split_asm Member.split_asm atype.split_asm) + then show ?thesis + proof cases + case (Contract c) + moreover from assms 1 2 3 4 Contract obtain ct cn f where 5: "ep $$ c = Some (ct, cn, f)" by (simp add: Let_def split: if_split_asm prod.split_asm result.split_asm Stackvalue.split_asm Type.split_asm Types.split_asm option.split_asm atype.split_asm atype.split_asm) + moreover from assms 1 2 3 4 Contract 5 obtain acc where 6: "transfer (address ev) adv v' (accounts st) = Some acc" by (simp add: Let_def split: if_split_asm prod.split_asm result.split_asm Stackvalue.split_asm Type.split_asm Types.split_asm option.split_asm Member.split_asm) + moreover from assms 1 2 3 4 Contract 5 6 obtain st'' where 7: "stmt f (e' ct adv c (st\gas := g'\) v') emptyStore (st\gas := g', accounts := acc, stack:=emptyStore, memory:=emptyStore\) = Normal ((), st'')" by (simp add: Let_def split: if_split_asm prod.split_asm result.split_asm Stackvalue.split_asm Type.split_asm Types.split_asm option.split_asm Member.split_asm) + moreover from assms 1 2 3 4 Contract 5 6 7 have "st' = st''\stack:=stack st, memory := memory st\" by (simp add: Let_def split: if_split_asm prod.split_asm result.split_asm Stackvalue.split_asm Type.split_asm Types.split_asm option.split_asm Member.split_asm) + ultimately show ?thesis using 1 2 3 4 that(1) by simp + next + case EOA + moreover from assms 1 2 3 4 EOA obtain acc where 5: "transfer (address ev) adv v' (accounts st) = Some acc" by (simp add: Let_def split: if_split_asm prod.split_asm result.split_asm Stackvalue.split_asm Type.split_asm Types.split_asm option.split_asm Member.split_asm) + moreover from assms 1 2 3 4 EOA 5 have "st' = st\gas:=g', accounts:=acc\" by (simp add: Let_def split: if_split_asm prod.split_asm result.split_asm Stackvalue.split_asm Type.split_asm Types.split_asm option.split_asm Member.split_asm) + ultimately show ?thesis using 1 2 3 4 that(2) by simp + qed +qed + +lemma blockNone: + fixes ev + assumes "stmt (BLOCK ((id0, tp), None) s) ev cd st = Normal (x, st')" + obtains cd' mem' sck' e' + where "gas st > costs (BLOCK ((id0, tp), None) s) ev cd st" + and "decl id0 tp None False cd (memory (st\gas := gas st - costs (BLOCK ((id0, tp), None) s) ev cd st\)) (storage (st\gas := gas st - costs (BLOCK ((id0, tp), None) s) ev cd st\)) (cd, memory (st\gas := gas st - costs (BLOCK ((id0, tp), None) s) ev cd st\), stack (st\gas := gas st - costs (BLOCK ((id0, tp), None) s) ev cd st\), ev) = Some (cd', mem', sck', e')" + and "stmt s e' cd' (st\gas := gas st - costs (BLOCK ((id0, tp), None) s) ev cd st, stack := sck', memory := mem'\) = Normal ((), st')" + using assms by (simp split:if_split_asm prod.split_asm option.split_asm) + +lemma blockSome: + fixes ev + assumes "stmt (BLOCK ((id0, tp), Some ex') s) ev cd st = Normal (x, st')" + obtains v t g cd' mem' sck' e' + where "gas st > costs (BLOCK ((id0, tp), Some ex') s) ev cd st" + and "expr ex' ev cd (st\gas := gas st - costs (BLOCK ((id0, tp), Some ex') s) ev cd st\) (gas st - costs (BLOCK ((id0, tp), Some ex') s) ev cd st) = Normal((v,t),g)" + and "decl id0 tp (Some (v, t)) False cd (memory (st\gas := g\)) (storage (st\gas := g\)) (cd, memory (st\gas := g\), stack (st\gas := g\), ev) = Some (cd', mem', sck', e')" + and "stmt s e' cd' (st\gas := g, stack := sck', memory := mem'\) = Normal ((), st')" + using assms by (auto split:if_split_asm result.split_asm prod.split_asm option.split_asm) + +lemma new: + fixes i xe val ev cd st + defines "st0 \ st\gas := gas st - costs (NEW i xe val) ev cd st\" + defines "adv0 \ hash (address ev) (ShowL\<^sub>n\<^sub>a\<^sub>t (contracts (accounts st0 (address ev))))" + defines "st1 g \ st\gas := g, accounts := (accounts st)(adv0 := \bal = ShowL\<^sub>i\<^sub>n\<^sub>t 0, type = Some (Contract i), contracts = 0\), storage:=(storage st)(adv0 := {$$})\" + defines "e' members c v \ ffold (init members) (emptyEnv adv0 c (address ev) v) (fmdom members)" + assumes "stmt (NEW i xe val) ev cd st = Normal (x, st')" + obtains v t g ct cn fb e\<^sub>l cd\<^sub>l k\<^sub>l m\<^sub>l g' acc st'' + where "gas st > costs (NEW i xe val) ev cd st" + and "type (accounts st adv0) = None" + and "expr val ev cd st0 (gas st0) = Normal((KValue v, Value t),g)" + and "ep $$ i = Some (ct, cn, fb)" + and "load True (fst cn) xe (e' ct i v) emptyStore emptyStore emptyStore ev cd (st0\gas := g\) g = Normal ((e\<^sub>l, cd\<^sub>l, k\<^sub>l, m\<^sub>l), g')" + and "transfer (address ev) adv0 v (accounts (st1 g')) = Some acc" + and "stmt (snd cn) e\<^sub>l cd\<^sub>l (st1 g'\accounts := acc, stack:=k\<^sub>l, memory:=m\<^sub>l\) = Normal ((), st'')" + and "st' = incrementAccountContracts (address ev) (st''\stack:=stack st, memory := memory st\)" +proof - + from assms have 1: "gas st > costs (NEW i xe val) ev cd st" by (simp split:if_split_asm) + moreover from st0_def assms 1 have 2: "type (accounts st adv0) = None" by (simp split: if_split_asm) + moreover from st0_def assms 1 2 obtain v t g where 3: "expr val ev cd st0 (gas st0) = Normal((KValue v, Value t),g)" by (simp split: prod.split_asm result.split_asm Stackvalue.split_asm Type.split_asm) + moreover from assms 1 st0_def 2 3 obtain ct cn fb where 4: "ep $$ i = Some(ct, cn, fb)" by (simp split: prod.split_asm result.split_asm option.split_asm) + moreover from st0_def adv0_def e'_def assms 1 2 3 4 obtain e\<^sub>l cd\<^sub>l k\<^sub>l m\<^sub>l g' where 5: "load True (fst cn) xe (e' ct i v) emptyStore emptyStore emptyStore ev cd (st0\gas := g\) g = Normal ((e\<^sub>l, cd\<^sub>l, k\<^sub>l, m\<^sub>l), g')" by (simp add:Let_def split: prod.split_asm result.split_asm option.split_asm) + moreover from st0_def adv0_def e'_def assms 1 2 3 4 5 obtain acc where 6: "transfer (address ev) adv0 v (accounts (st1 g')) = Some acc" by (simp add:Let_def split: prod.split_asm result.split_asm option.split_asm) + moreover from st0_def st1_def adv0_def e'_def assms 1 2 3 4 5 6 obtain st'' where "stmt (snd cn) e\<^sub>l cd\<^sub>l (st1 g'\accounts := acc, stack:=k\<^sub>l, memory:=m\<^sub>l\) = Normal ((), st'')" by (simp add:Let_def split: prod.split_asm result.split_asm option.split_asm) + ultimately show ?thesis using that assms by simp +qed + +lemma atype_same: + assumes "stmt stm ev cd st = Normal (x, st')" + and "type (accounts st ad) = Some ctype" + shows "type (accounts st' ad) = Some ctype" +using assms +proof (induction arbitrary: st' rule: stmt.induct) + case (1 e cd st) + then show ?case using skip[OF 1(1)] by auto +next + case (2 lv ex env cd st) + show ?case by (cases rule: assign[OF 2(1)]; simp add: 2(2)) +next + case (3 s1 s2 e cd st) + show ?case + proof (cases rule: comp[OF 3(3)]) + case (1 st'') + then show ?thesis using 3 by simp + qed +next + case (4 ex s1 s2 e cd st) + show ?case + proof (cases rule: ite[OF 4(3)]) + case (1 g) + then show ?thesis using 4 by simp + next + case (2 g) + then show ?thesis using 4 by (simp split: if_split_asm) + qed +next + case (5 ex s0 e cd st) + show ?case + proof (cases rule: while[OF 5(3)]) + case (1 g st'') + then show ?thesis using 5 by simp + next + case (2 g) + then show ?thesis using 5 by simp + qed +next + case (6 i xe e cd st) + show ?case + proof (cases rule: invoke[OF 6(2)]) + case (1 ct fb fp f e\<^sub>l cd\<^sub>l k\<^sub>l m\<^sub>l g st'') + then show ?thesis using 6 by simp + qed +next + case (7 ad' i xe val e cd st) + show ?case + proof (cases rule: external[OF 7(3)]) + case (1 adv c g ct cn fb' v t g' v' fp f e\<^sub>l cd\<^sub>l k\<^sub>l m\<^sub>l g'' acc st'') + moreover from 7(4) have "type (acc ad) = Some ctype" using transfer_type_same[OF 1(10)] by simp + ultimately show ?thesis using 7(1) by simp + next + case (2 adv c g ct cn fb' v t g' v' acc st'') + moreover from 7(4) have "type (acc ad) = Some ctype" using transfer_type_same[OF 2(9)] by simp + ultimately show ?thesis using 7(2) by simp + qed +next + case (8 ad' ex e cd st) + show ?case + proof (cases rule: transfer[OF 8(2)]) + case (1 v t g adv c g' v' acc ct cn f st'') + moreover from 8(3) have "type (acc ad) = Some ctype" using transfer_type_same[OF 1(7)] by simp + ultimately show ?thesis using 8(1) by simp + next + case (2 v t g adv g' v' acc) + moreover from 8(3) have "type (acc ad) = Some ctype" using transfer_type_same[OF 2(6)] by simp + ultimately show ?thesis by simp + qed +next + case (9 id0 tp s e\<^sub>v cd st) + show ?case + proof (cases rule: blockNone[OF 9(2)]) + case (1 cd' mem' sck' e') + then show ?thesis using 9 by simp + qed +next + case (10 id0 tp ex' s e\<^sub>v cd st) + show ?case + proof (cases rule: blockSome[OF 10(2)]) + case (1 v t g cd' mem' sck' e') + then show ?thesis using 10 by simp + qed +next + case (11 i xe val e cd st) + show ?case + proof (cases rule: new[OF 11(2)]) + case (1 v t g ct cn fb e\<^sub>l cd\<^sub>l k\<^sub>l m\<^sub>l g' acc st'') + moreover have "hash (address e) \contracts (accounts st (address e))\ \ ad" using 11(3) 1(2) by auto + ultimately show ?thesis + using 11 transfer_type_same[OF 1(6)] incrementAccountContracts_type by simp + qed +qed + +declare lexp.simps[simp del] +declare stmt.simps[simp del] + +end + +subsection \A minimal cost model\ + +fun costs_min :: "S \ Environment \ CalldataT \ State \ Gas" + where + "costs_min SKIP e cd st = 0" +| "costs_min (ASSIGN lv ex) e cd st = 0" +| "costs_min (COMP s1 s2) e cd st = 0" +| "costs_min (ITE ex s1 s2) e cd st = 0" +| "costs_min (WHILE ex s0) e cd st = 1" +| "costs_min (TRANSFER ad ex) e cd st = 1" +| "costs_min (BLOCK ((id0, tp), ex) s) e cd st =0" +| "costs_min (INVOKE _ _) e cd st = 1" +| "costs_min (EXTERNAL _ _ _ _) e cd st = 1" +| "costs_min (NEW _ _ _) e cd st = 1" + +fun costs_ex :: "E \ Environment \ CalldataT \ State \ Gas" + where + "costs_ex (E.INT _ _) e cd st = 0" +| "costs_ex (UINT _ _) e cd st = 0" +| "costs_ex (ADDRESS _) e cd st = 0" +| "costs_ex (BALANCE _) e cd st = 0" +| "costs_ex THIS e cd st = 0" +| "costs_ex SENDER e cd st = 0" +| "costs_ex VALUE e cd st = 0" +| "costs_ex (TRUE) e cd st = 0" +| "costs_ex (FALSE) e cd st = 0" +| "costs_ex (LVAL _) e cd st = 0" +| "costs_ex (PLUS _ _) e cd st = 0" +| "costs_ex (MINUS _ _) e cd st = 0" +| "costs_ex (EQUAL _ _) e cd st = 0" +| "costs_ex (LESS _ _) e cd st = 0" +| "costs_ex (AND _ _) e cd st = 0" +| "costs_ex (OR _ _) e cd st = 0" +| "costs_ex (NOT _) e cd st = 0" +| "costs_ex (CALL _ _) e cd st = 1" +| "costs_ex (ECALL _ _ _) e cd st = 1" +| "costs_ex CONTRACTS e cd st = 0" + +global_interpretation solidity: statement_with_gas costs_ex fmempty costs_min + defines stmt = "solidity.stmt" + and lexp = solidity.lexp + and expr = solidity.expr + and ssel = solidity.ssel + and rexp = solidity.rexp + and msel = solidity.msel + and load = solidity.load + by unfold_locales auto + +section \Examples\ + +subsection \msel\ + +abbreviation mymemory2::MemoryT + where "mymemory2 \ + \mapping = fmap_of_list + [(STR ''3.2'', MPointer STR ''5'')], + toploc = 1\" + +(* + FIXME + +lemma "msel True (MTArray 5 (MTArray 6 (MTValue TBool))) (STR ''2'') [UINT 8 3] eempty emptyStore (mystate\gas:=1\) 1 += Normal ((STR ''3.2'', MTArray 6 (MTValue TBool)), 1)" by Solidity_Symbex.solidity_symbex + +lemma "msel True (MTArray 5 (MTArray 6 (MTValue TBool))) (STR ''2'') [UINT 8 3, UINT 8 4] eempty emptyStore (mystate\gas:=1,memory:=mymemory2\) 1 += Normal ((STR ''4.5'', MTValue TBool), 2)" by Solidity_Symbex.solidity_symbex +*) + +lemma "msel True (MTArray 5 (MTArray 6 (MTValue TBool))) (STR ''2'') [UINT 8 5] eempty emptyStore (mystate\gas:=1,memory:=mymemory2\) 1 += Exception (Err)" by Solidity_Symbex.solidity_symbex + +end diff --git a/thys/Solidity/Storage.thy b/thys/Solidity/Storage.thy --- a/thys/Solidity/Storage.thy +++ b/thys/Solidity/Storage.thy @@ -1,240 +1,301 @@ -section\Storage\ - -theory Storage -imports Valuetypes "HOL-Library.Finite_Map" - -begin - -(*Covered*) -fun hash :: "Location \ String.literal \ Location" -where "hash loc ix = ix + (STR ''.'' + loc)" - -subsection \General Store\ - -(*Covered*) -record 'v Store = - mapping :: "(Location,'v) fmap" - toploc :: nat - -fun accessStore :: "Location \ 'v Store \ 'v option" -where "accessStore loc st = fmlookup (mapping st) loc" - -definition emptyStore :: "'v Store" -where "emptyStore = \ mapping=fmempty, toploc=0 \" - -declare emptyStore_def [solidity_symbex] - -fun allocate :: "'v Store \ Location * ('v Store)" -where "allocate s = (let ntop = Suc(toploc s) in (ShowL\<^sub>n\<^sub>a\<^sub>t ntop, s \toploc := ntop\))" - -fun updateStore :: "Location \ 'v \ 'v Store \ 'v Store" -where "updateStore loc val s = s \ mapping := fmupd loc val (mapping s)\" - -fun push :: "'v \ 'v Store \ 'v Store" - where "push val sto = (let s = updateStore (ShowL\<^sub>n\<^sub>a\<^sub>t (toploc sto)) val sto in snd (allocate s))" - -subsection \Stack\ -(*Covered*) -datatype Stackvalue = KValue Valuetype - | KCDptr Location - | KMemptr Location - | KStoptr Location - -(*Covered*) -type_synonym Stack = "Stackvalue Store" - -subsection \Storage\ - -subsubsection \Definition\ - -type_synonym Storagevalue = Valuetype - -(*Covered*) -type_synonym StorageT = "(Location,Storagevalue) fmap" - -(*Covered*) -datatype STypes = STArray int STypes - | STMap Types STypes - | STValue Types - -subsubsection \Example\ - -abbreviation mystorage::StorageT -where "mystorage \ (fmap_of_list - [(STR ''0.0.1'', STR ''True''), - (STR ''1.0.1'', STR ''False''), - (STR ''0.1.1'', STR ''True''), - (STR ''1.1.1'', STR ''False'')])" - -subsubsection \Access storage\ - -(*Covered*) -fun accessStorage :: "Types \ Location \ StorageT \ Storagevalue" -where - "accessStorage t loc sto = - (case fmlookup sto loc of - Some v \ v - | None \ ival t)" - -subsubsection \Copy from storage to storage\ - -fun copyRec :: "Location \ Location \ STypes \ StorageT \ StorageT option" -where - "copyRec loc loc' (STArray x t) sto = - iter' (\i s'. copyRec (hash loc (ShowL\<^sub>i\<^sub>n\<^sub>t i)) (hash loc' (ShowL\<^sub>i\<^sub>n\<^sub>t i)) t s') sto x" -| "copyRec loc loc' (STValue t) sto = - (let e = accessStorage t loc sto in Some (fmupd loc' e sto))" -| "copyRec _ _ (STMap _ _) _ = None" - -fun copy :: "Location \ Location \ int \ STypes \ StorageT \ StorageT option" -where - "copy loc loc' x t sto = - iter' (\i s'. copyRec (hash loc (ShowL\<^sub>i\<^sub>n\<^sub>t i)) (hash loc' (ShowL\<^sub>i\<^sub>n\<^sub>t i)) t s') sto x" - -subsection \Memory and Calldata\ - -subsubsection \Definition\ - -(*Covered*) -datatype Memoryvalue = - MValue Valuetype - | MPointer Location -(*Covered*) -type_synonym MemoryT = "Memoryvalue Store" -(*Covered*) -type_synonym CalldataT = MemoryT -(*Covered*) -datatype MTypes = MTArray int MTypes - | MTValue Types - -subsubsection \Example\ - -abbreviation mymemory::MemoryT - where "mymemory \ - \mapping = fmap_of_list - [(STR ''1.1.0'', MValue STR ''False''), - (STR ''0.1.0'', MValue STR ''True''), - (STR ''1.0'', MPointer STR ''1.0''), - (STR ''1.0.0'', MValue STR ''False''), - (STR ''0.0.0'', MValue STR ''True''), - (STR ''0.0'', MPointer STR ''0.0'')], - toploc = 1\" - -subsubsection \Initialization\ - -subsubsection \Definition\ - -(*Covered*) -fun minitRec :: "Location \ MTypes \ MemoryT \ MemoryT" -where - "minitRec loc (MTArray x t) = (\mem. - let m = updateStore loc (MPointer loc) mem - in iter (\i m' . minitRec (hash loc (ShowL\<^sub>i\<^sub>n\<^sub>t i)) t m') m x)" -| "minitRec loc (MTValue t) = updateStore loc (MValue (ival t))" - -(*Covered*) -fun minit :: "int \ MTypes \ MemoryT \ MemoryT" -where - "minit x t mem = - (let l = ShowL\<^sub>n\<^sub>a\<^sub>t (toploc mem); - m = iter (\i m' . minitRec (hash l (ShowL\<^sub>i\<^sub>n\<^sub>t i)) t m') mem x - in snd (allocate m))" - -subsubsection \Example\ - -lemma "minit 2 (MTArray 2 (MTValue TBool)) emptyStore = -\mapping = fmap_of_list - [(STR ''0.0'', MPointer STR ''0.0''), (STR ''0.0.0'', MValue STR ''False''), - (STR ''1.0.0'', MValue STR ''False''), (STR ''1.0'', MPointer STR ''1.0''), - (STR ''0.1.0'', MValue STR ''False''), (STR ''1.1.0'', MValue STR ''False'')], - toploc = 1\" by eval - -subsubsection \Copy from memory to memory\ - -subsubsection \Definition\ - -fun cpm2mrec :: "Location \ Location \ MTypes \ MemoryT \ MemoryT \ MemoryT option" -where - "cpm2mrec l\<^sub>s l\<^sub>d (MTArray x t) m\<^sub>s m\<^sub>d = - (case accessStore l\<^sub>s m\<^sub>s of - Some e \ - (case e of - MPointer l \ (let m = updateStore l\<^sub>d (MPointer l\<^sub>d) m\<^sub>d - in iter' (\i m'. cpm2mrec (hash l\<^sub>s (ShowL\<^sub>i\<^sub>n\<^sub>t i)) (hash l\<^sub>d (ShowL\<^sub>i\<^sub>n\<^sub>t i)) t m\<^sub>s m') m x) - | _ \ None) - | None \ None)" -| "cpm2mrec l\<^sub>s l\<^sub>d (MTValue t) m\<^sub>s m\<^sub>d = - (case accessStore l\<^sub>s m\<^sub>s of - Some e \ (case e of - MValue v \ Some (updateStore l\<^sub>d (MValue v) m\<^sub>d) - | _ \ None) - | None \ None)" - -fun cpm2m :: "Location \ Location \ int \ MTypes \ MemoryT \ MemoryT \ MemoryT option" -where - "cpm2m l\<^sub>s l\<^sub>d x t m\<^sub>s m\<^sub>d = iter' (\i m. cpm2mrec (hash l\<^sub>s (ShowL\<^sub>i\<^sub>n\<^sub>t i)) (hash l\<^sub>d (ShowL\<^sub>i\<^sub>n\<^sub>t i)) t m\<^sub>s m) m\<^sub>d x" - -subsubsection \Example\ - -lemma "cpm2m (STR ''0'') (STR ''0'') 2 (MTArray 2 (MTValue TBool)) mymemory (snd (allocate emptyStore)) = Some mymemory" - by eval - -subsection \Copy from storage to memory\ - -subsubsection \Definition\ - -(*Covered*) -fun cps2mrec :: "Location \ Location \ STypes \ StorageT \ MemoryT \ MemoryT option" -where - "cps2mrec locs locm (STArray x t) sto mem = - (let m = updateStore locm (MPointer locm) mem - in iter' (\i m'. cps2mrec (hash locs (ShowL\<^sub>i\<^sub>n\<^sub>t i)) (hash locm (ShowL\<^sub>i\<^sub>n\<^sub>t i)) t sto m') m x)" -| "cps2mrec locs locm (STValue t) sto mem = - (let v = accessStorage t locs sto - in Some (updateStore locm (MValue v) mem))" -| "cps2mrec _ _ (STMap _ _) _ _ = None" - -(*Covered*) -fun cps2m :: "Location \ Location \ int \ STypes \ StorageT \ MemoryT \ MemoryT option" -where - "cps2m locs locm x t sto mem = - iter' (\i m'. cps2mrec (hash locs (ShowL\<^sub>i\<^sub>n\<^sub>t i)) (hash locm (ShowL\<^sub>i\<^sub>n\<^sub>t i)) t sto m') mem x" - -subsubsection \Example\ - -lemma "cps2m (STR ''1'') (STR ''0'') 2 (STArray 2 (STValue TBool)) mystorage (snd (allocate emptyStore)) = Some mymemory" - by eval - -subsection \Copy from memory to storage\ - -subsubsection \Definition\ - -(*covered*) -fun cpm2srec :: "Location \ Location \ MTypes \ MemoryT \ StorageT \ StorageT option" -where - "cpm2srec locm locs (MTArray x t) mem sto = - (case accessStore locm mem of - Some e \ - (case e of - MPointer l \ iter' (\i s'. cpm2srec (hash locm (ShowL\<^sub>i\<^sub>n\<^sub>t i)) (hash locs (ShowL\<^sub>i\<^sub>n\<^sub>t i)) t mem s') sto x - | _ \ None) - | None \ None)" -| "cpm2srec locm locs (MTValue t) mem sto = - (case accessStore locm mem of - Some e \ (case e of - MValue v \ Some (fmupd locs v sto) - | _ \ None) - | None \ None)" - -(*covered*) -fun cpm2s :: "Location \ Location \ int \ MTypes \ MemoryT \ StorageT \ StorageT option" -where - "cpm2s locm locs x t mem sto = - iter' (\i s'. cpm2srec (hash locm (ShowL\<^sub>i\<^sub>n\<^sub>t i)) (hash locs (ShowL\<^sub>i\<^sub>n\<^sub>t i)) t mem s') sto x" - -subsubsection \Example\ - -lemma "cpm2s (STR ''0'') (STR ''1'') 2 (MTArray 2 (MTValue TBool)) mymemory fmempty = Some mystorage" - by eval - -end +section\Storage\ + +theory Storage +imports Valuetypes "Finite-Map-Extras.Finite_Map_Extras" + +begin + +subsection \Hashing\ + +definition hash :: "Location \ String.literal \ Location" + where "hash loc ix = ix + (STR ''.'' + loc)" + +lemma example: "hash (STR ''1.0'') (STR ''2'') = hash (STR ''0'') (STR ''2.1'')" by eval + +lemma hash_explode: + "String.explode (hash l i) = String.explode i @ (String.explode (STR ''.'') @ String.explode l)" + unfolding hash_def by (simp add: plus_literal.rep_eq) + +lemma hash_dot: + "String.explode (hash l i) ! length (String.explode i) = CHR ''.''" + unfolding hash_def by (simp add: Literal.rep_eq plus_literal.rep_eq) + +lemma hash_injective: + assumes "hash l i = hash l' i'" + and "CHR ''.'' \ set (String.explode i)" + and "CHR ''.'' \ set (String.explode i')" + shows "l = l' \ i = i'" +proof (rule ccontr) + assume "\ (l = l' \ i = i')" + then consider (1) "i\i'" | (2) "i=i' \ l\l'" by auto + then have "String.explode (hash l i) \ String.explode (hash l' i')" + proof cases + case 1 + then have neq: "(String.explode i) \ (String.explode i')" by (metis literal.explode_inverse) + + consider (1) "length (String.explode i) = length (String.explode i')" | (2) "length (String.explode i) < length (String.explode i')" | (3) "length (String.explode i) > length (String.explode i')" by linarith + then show ?thesis + proof (cases) + case 1 + then obtain j where "String.explode i ! j \ String.explode i' ! j" using neq nth_equalityI by auto + then show ?thesis using 1 plus_literal.rep_eq unfolding hash_def by force + next + case 2 + then have "String.explode i' ! length (String.explode i) \ CHR ''.''" using assms(3) by (metis nth_mem) + then have "String.explode (hash l' i') ! length (String.explode i) \ CHR ''.''" using 2 hash_explode[of l' i'] by (simp add: nth_append) + moreover have "String.explode (hash l i) ! length (String.explode i) = CHR ''.''" using hash_dot by simp + ultimately show ?thesis by auto + next + case 3 + then have "String.explode i ! length (String.explode i') \ CHR ''.''" using assms(2) by (metis nth_mem) + then have "String.explode (hash l i) ! length (String.explode i') \ CHR ''.''" using 3 hash_explode[of l i] by (simp add: nth_append) + moreover have "String.explode (hash l' i') ! length (String.explode i') = CHR ''.''" using hash_dot by simp + ultimately show ?thesis by auto + qed + next + case 2 + then show ?thesis using hash_explode literal.explode_inject by force + qed + then show False using assms(1) by simp +qed + +subsection \General Store\ + +record 'v Store = + mapping :: "(Location,'v) fmap" + toploc :: nat + +definition accessStore :: "Location \ 'v Store \ 'v option" +where "accessStore loc st = fmlookup (mapping st) loc" + +definition emptyStore :: "'v Store" +where "emptyStore = \ mapping=fmempty, toploc=0 \" + +declare emptyStore_def [solidity_symbex] + +definition allocate :: "'v Store \ Location * ('v Store)" +where "allocate s = (let ntop = Suc(toploc s) in (ShowL\<^sub>n\<^sub>a\<^sub>t ntop, s \toploc := ntop\))" + +definition updateStore :: "Location \ 'v \ 'v Store \ 'v Store" +where "updateStore loc val s = s \ mapping := fmupd loc val (mapping s)\" + +definition push :: "'v \ 'v Store \ 'v Store" + where "push val sto = (let s = updateStore (ShowL\<^sub>n\<^sub>a\<^sub>t (toploc sto)) val sto in snd (allocate s))" + +subsection \Stack\ + +datatype Stackvalue = KValue Valuetype + | KCDptr Location + | KMemptr Location + | KStoptr Location + +type_synonym Stack = "Stackvalue Store" + +subsection \Storage\ + +subsubsection \Definition\ + +type_synonym Storagevalue = Valuetype + +type_synonym StorageT = "(Location,Storagevalue) fmap" + +datatype STypes = STArray int STypes + | STMap Types STypes + | STValue Types + +subsubsection \Example\ + +abbreviation mystorage::StorageT +where "mystorage \ (fmap_of_list + [(STR ''0.0.0'', STR ''False''), + (STR ''1.1.0'', STR ''True'')])" + +subsubsection \Access storage\ + +definition accessStorage :: "Types \ Location \ StorageT \ Storagevalue" +where + "accessStorage t loc sto = + (case sto $$ loc of + Some v \ v + | None \ ival t)" + +subsubsection \Copy from storage to storage\ + +primrec copyRec :: "Location \ Location \ STypes \ StorageT \ StorageT option" +where + "copyRec l\<^sub>s l\<^sub>d (STArray x t) sto = + iter' (\i s'. copyRec (hash l\<^sub>s (ShowL\<^sub>i\<^sub>n\<^sub>t i)) (hash l\<^sub>d (ShowL\<^sub>i\<^sub>n\<^sub>t i)) t s') sto x" +| "copyRec l\<^sub>s l\<^sub>d (STValue t) sto = + (let e = accessStorage t l\<^sub>s sto in Some (fmupd l\<^sub>d e sto))" +| "copyRec _ _ (STMap _ _) _ = None" + +definition copy :: "Location \ Location \ int \ STypes \ StorageT \ StorageT option" +where + "copy l\<^sub>s l\<^sub>d x t sto = + iter' (\i s'. copyRec (hash l\<^sub>s (ShowL\<^sub>i\<^sub>n\<^sub>t i)) (hash l\<^sub>d (ShowL\<^sub>i\<^sub>n\<^sub>t i)) t s') sto x" + +abbreviation mystorage2::StorageT +where "mystorage2 \ (fmap_of_list + [(STR ''0.0.0'', STR ''False''), + (STR ''1.1.0'', STR ''True''), + (STR ''0.5'', STR ''False''), + (STR ''1.5'', STR ''True'')])" + +lemma "copy (STR ''1.0'') (STR ''5'') 2 (STValue TBool) mystorage = Some mystorage2" + by eval + +subsection \Memory and Calldata\ + +subsubsection \Definition\ + +datatype Memoryvalue = + MValue Valuetype + | MPointer Location + +type_synonym MemoryT = "Memoryvalue Store" + +type_synonym CalldataT = MemoryT + +datatype MTypes = + MTArray int MTypes + | MTValue Types + +subsubsection \Example\ + +abbreviation mymemory::MemoryT + where "mymemory \ + \mapping = fmap_of_list + [(STR ''1.1.0'', MValue STR ''False''), + (STR ''0.1.0'', MValue STR ''True''), + (STR ''1.0'', MPointer STR ''1.0''), + (STR ''1.0.0'', MValue STR ''False''), + (STR ''0.0.0'', MValue STR ''True''), + (STR ''0.0'', MPointer STR ''0.0'')], + toploc = 1\" + +subsubsection \Initialization\ + +subsubsection \Definition\ + +primrec minitRec :: "Location \ MTypes \ MemoryT \ MemoryT" +where + "minitRec loc (MTArray x t) = (\mem. + let m = updateStore loc (MPointer loc) mem + in iter (\i m' . minitRec (hash loc (ShowL\<^sub>i\<^sub>n\<^sub>t i)) t m') m x)" +| "minitRec loc (MTValue t) = updateStore loc (MValue (ival t))" + +definition minit :: "int \ MTypes \ MemoryT \ MemoryT" +where + "minit x t mem = + (let l = ShowL\<^sub>n\<^sub>a\<^sub>t (toploc mem); + m = iter (\i m' . minitRec (hash l (ShowL\<^sub>i\<^sub>n\<^sub>t i)) t m') mem x + in snd (allocate m))" + +subsubsection \Example\ + +lemma "minit 2 (MTArray 2 (MTValue TBool)) emptyStore = +\mapping = fmap_of_list + [(STR ''0.0'', MPointer STR ''0.0''), (STR ''0.0.0'', MValue STR ''False''), + (STR ''1.0.0'', MValue STR ''False''), (STR ''1.0'', MPointer STR ''1.0''), + (STR ''0.1.0'', MValue STR ''False''), (STR ''1.1.0'', MValue STR ''False'')], + toploc = 1\" by eval + +subsubsection \Copy from memory to memory\ + +subsubsection \Definition\ + +primrec cpm2mrec :: "Location \ Location \ MTypes \ MemoryT \ MemoryT \ MemoryT option" +where + "cpm2mrec l\<^sub>s l\<^sub>d (MTArray x t) m\<^sub>s m\<^sub>d = + (case accessStore l\<^sub>s m\<^sub>s of + Some (MPointer l) \ + (let m = updateStore l\<^sub>d (MPointer l\<^sub>d) m\<^sub>d + in iter' (\i m'. cpm2mrec (hash l\<^sub>s (ShowL\<^sub>i\<^sub>n\<^sub>t i)) (hash l\<^sub>d (ShowL\<^sub>i\<^sub>n\<^sub>t i)) t m\<^sub>s m') m x) + | _ \ None)" +| "cpm2mrec l\<^sub>s l\<^sub>d (MTValue t) m\<^sub>s m\<^sub>d = + (case accessStore l\<^sub>s m\<^sub>s of + Some (MValue v) \ Some (updateStore l\<^sub>d (MValue v) m\<^sub>d) + | _ \ None)" + +definition cpm2m :: "Location \ Location \ int \ MTypes \ MemoryT \ MemoryT \ MemoryT option" +where + "cpm2m l\<^sub>s l\<^sub>d x t m\<^sub>s m\<^sub>d = iter' (\i m. cpm2mrec (hash l\<^sub>s (ShowL\<^sub>i\<^sub>n\<^sub>t i)) (hash l\<^sub>d (ShowL\<^sub>i\<^sub>n\<^sub>t i)) t m\<^sub>s m) m\<^sub>d x" + +subsubsection \Example\ + +lemma "cpm2m (STR ''0'') (STR ''0'') 2 (MTArray 2 (MTValue TBool)) mymemory (snd (allocate emptyStore)) = Some mymemory" + by eval + +abbreviation mymemory2::MemoryT + where "mymemory2 \ + \mapping = fmap_of_list + [(STR ''0.5'', MValue STR ''True''), + (STR ''1.5'', MValue STR ''False'')], + toploc = 0\" + +lemma "cpm2m (STR ''1.0'') (STR ''5'') 2 (MTValue TBool) mymemory emptyStore = Some mymemory2" by eval + +subsection \Copy from storage to memory\ + +subsubsection \Definition\ + +primrec cps2mrec :: "Location \ Location \ STypes \ StorageT \ MemoryT \ MemoryT option" +where + "cps2mrec locs locm (STArray x t) sto mem = + (let m = updateStore locm (MPointer locm) mem + in iter' (\i m'. cps2mrec (hash locs (ShowL\<^sub>i\<^sub>n\<^sub>t i)) (hash locm (ShowL\<^sub>i\<^sub>n\<^sub>t i)) t sto m') m x)" +| "cps2mrec locs locm (STValue t) sto mem = + (let v = accessStorage t locs sto + in Some (updateStore locm (MValue v) mem))" +| "cps2mrec _ _ (STMap _ _) _ _ = None" + +definition cps2m :: "Location \ Location \ int \ STypes \ StorageT \ MemoryT \ MemoryT option" +where + "cps2m locs locm x t sto mem = + iter' (\i m'. cps2mrec (hash locs (ShowL\<^sub>i\<^sub>n\<^sub>t i)) (hash locm (ShowL\<^sub>i\<^sub>n\<^sub>t i)) t sto m') mem x" + +subsubsection \Example\ + +abbreviation mystorage3::StorageT +where "mystorage3 \ (fmap_of_list + [(STR ''0.0.1'', STR ''True''), + (STR ''1.0.1'', STR ''False''), + (STR ''0.1.1'', STR ''True''), + (STR ''1.1.1'', STR ''False'')])" + +lemma "cps2m (STR ''1'') (STR ''0'') 2 (STArray 2 (STValue TBool)) mystorage3 (snd (allocate emptyStore)) = Some mymemory" + by eval + +subsection \Copy from memory to storage\ + +subsubsection \Definition\ + +primrec cpm2srec :: "Location \ Location \ MTypes \ MemoryT \ StorageT \ StorageT option" +where + "cpm2srec locm locs (MTArray x t) mem sto = + (case accessStore locm mem of + Some (MPointer l) \ + iter' (\i s'. cpm2srec (hash locm (ShowL\<^sub>i\<^sub>n\<^sub>t i)) (hash locs (ShowL\<^sub>i\<^sub>n\<^sub>t i)) t mem s') sto x + | _ \ None)" +| "cpm2srec locm locs (MTValue t) mem sto = + (case accessStore locm mem of + Some (MValue v) \ Some (fmupd locs v sto) + | _ \ None)" + +definition cpm2s :: "Location \ Location \ int \ MTypes \ MemoryT \ StorageT \ StorageT option" +where + "cpm2s locm locs x t mem sto = + iter' (\i s'. cpm2srec (hash locm (ShowL\<^sub>i\<^sub>n\<^sub>t i)) (hash locs (ShowL\<^sub>i\<^sub>n\<^sub>t i)) t mem s') sto x" + +subsubsection \Example\ + +lemma "cpm2s (STR ''0'') (STR ''1'') 2 (MTArray 2 (MTValue TBool)) mymemory fmempty = Some mystorage3" + by eval + +declare copyRec.simps [simp del] +declare minitRec.simps [simp del] +declare cpm2mrec.simps [simp del] +declare cps2mrec.simps [simp del] +declare cpm2srec.simps [simp del] + +end diff --git a/thys/Solidity/Utils.thy b/thys/Solidity/Utils.thy new file mode 100644 --- /dev/null +++ b/thys/Solidity/Utils.thy @@ -0,0 +1,17 @@ +theory Utils +imports + Main + "HOL-Eisbach.Eisbach" +begin + +method solve methods m = (m ; fail) + +named_theorems intros +declare conjI[intros] impI[intros] allI[intros] +method intros = (rule intros; intros?) + +named_theorems elims +method elims = ((rule intros | erule elims); elims?) +declare conjE[elims] + +end \ No newline at end of file diff --git a/thys/Solidity/Valuetypes.thy b/thys/Solidity/Valuetypes.thy --- a/thys/Solidity/Valuetypes.thy +++ b/thys/Solidity/Valuetypes.thy @@ -1,252 +1,278 @@ -section\Value Types\ - -theory Valuetypes -imports ReadShow -begin - -fun iter :: "(int \ 'b \ 'b) \ 'b \ int \ 'b" -where - "iter f v x = (if x \ 0 then v - else f (x-1) (iter f v (x-1)))" - -fun iter' :: "(int \ 'b \ 'b option) \ 'b \ int \ 'b option" -where - "iter' f v x = (if x \ 0 then Some v - else case iter' f v (x-1) of - Some v' \ f (x-1) v' - | None \ None)" - -type_synonym Address = String.literal -type_synonym Location = String.literal -type_synonym Valuetype = String.literal - -(*Covered*) -datatype Types = TSInt nat - | TUInt nat - | TBool - | TAddr - -(*Covered*) -fun createSInt :: "nat \ int \ Valuetype" -where - "createSInt b v = - (if v \ 0 - then ShowL\<^sub>i\<^sub>n\<^sub>t (-(2^(b-1)) + (v+2^(b-1)) mod (2^b)) - else ShowL\<^sub>i\<^sub>n\<^sub>t (2^(b-1) - (-v+2^(b-1)-1) mod (2^b) - 1))" - -lemma upper_bound: - fixes b::nat - and c::int - assumes "b > 0" - and "c < 2^(b-1)" - shows "c + 2^(b-1) < 2^b" -proof - - have a1: "\P. (\b::nat. P b) \ (\b>0. P ((b-1)::nat))" by simp - have b2: "\b::nat. (\(c::int)<2^b. (c + 2^b) < 2^(Suc b))" by simp - show ?thesis using a1[OF b2] assms by simp -qed - -lemma upper_bound2: - fixes b::nat - and c::int - assumes "b > 0" - and "c < 2^b" - and "c \ 0" - shows "c - (2^(b-1)) < 2^(b-1)" -proof - - have a1: "\P. (\b::nat. P b) \ (\b>0. P ((b-1)::nat))" by simp - have b2: "\b::nat. (\(c::int)<2^(Suc b). c\0 \ (c - 2^b) < 2^b)" by simp - show ?thesis using a1[OF b2] assms by simp -qed - -lemma upper_bound3: - fixes b::nat - and v::int - defines "x \ - (2 ^ (b - 1)) + (v + 2 ^ (b - 1)) mod 2 ^ b" - assumes "b>0" - shows "x < 2^(b-1)" - using upper_bound2 assms by auto - -lemma lower_bound: - fixes b::nat - assumes "b>0" - shows "\(c::int) \ -(2^(b-1)). (-c + 2^(b-1) - 1 < 2^b)" -proof - - have a1: "\P. (\b::nat. P b) \ (\b>0. P ((b-1)::nat))" by simp - have b2: "\b::nat. \(c::int) \ -(2^b). (-c + (2^b) - 1) < 2^(Suc b)" by simp - show ?thesis using a1[OF b2] assms by simp -qed - -lemma lower_bound2: - fixes b::nat - and v::int - defines "x \ 2^(b - 1) - (-v+2^(b-1)-1) mod 2^b - 1" - assumes "b>0" - shows "x \ - (2 ^ (b - 1))" - using upper_bound2 assms by auto - -lemma createSInt_id_g0: - fixes b::nat - and v::int - assumes "v \ 0" - and "v < 2^(b-1)" - and "b > 0" - shows "createSInt b v = ShowL\<^sub>i\<^sub>n\<^sub>t v" -proof - - from assms have "v + 2^(b-1) \ 0" by simp - moreover from assms have "v + (2^(b-1)) < 2^b" using upper_bound[of b] by auto - ultimately have "(v + 2^(b-1)) mod (2^b) = v + 2^(b-1)" by simp - moreover from assms have "createSInt b v=ShowL\<^sub>i\<^sub>n\<^sub>t (-(2^(b-1)) + (v+2^(b-1)) mod (2^b))" by simp - ultimately show ?thesis by simp -qed - -lemma createSInt_id_l0: - fixes b::nat - and v::int - assumes "v < 0" - and "v \ -(2^(b-1))" - and "b > 0" - shows "createSInt b v = ShowL\<^sub>i\<^sub>n\<^sub>t v" -proof - - from assms have "-v + 2^(b-1) - 1 \ 0" by simp - moreover from assms have "-v + 2^(b-1) - 1 < 2^b" using lower_bound[of b] by auto - ultimately have "(-v + 2^(b-1) - 1) mod (2^b) = (-v + 2^(b-1) - 1)" by simp - moreover from assms have "createSInt b v= ShowL\<^sub>i\<^sub>n\<^sub>t (2^(b-1) - (-v+2^(b-1)-1) mod (2^b) - 1)" by simp - ultimately show ?thesis by simp -qed - -lemma createSInt_id: - fixes b::nat - and v::int - assumes "v < 2^(b-1)" - and "v \ -(2^(b-1))" - and "b > 0" - shows "createSInt b v = ShowL\<^sub>i\<^sub>n\<^sub>t v" using createSInt_id_g0 createSInt_id_l0 assms by simp - -(*Covered*) -fun createUInt :: "nat \ int \ Valuetype" - where "createUInt b v = ShowL\<^sub>i\<^sub>n\<^sub>t (v mod (2^b))" - -lemma createUInt_id: - assumes "v \ 0" - and "v < 2^b" - shows "createUInt b v = ShowL\<^sub>i\<^sub>n\<^sub>t v" -by (simp add: assms(1) assms(2)) - -fun createBool :: "bool \ Valuetype" -where - "createBool b = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l b" - -fun createAddress :: "Address \ Valuetype" -where - "createAddress ad = ad" - -fun convert :: "Types \ Types \ Valuetype \ (Valuetype * Types) option" -where - "convert (TSInt b1) (TSInt b2) v = - (if b1 \ b2 - then Some (v, TSInt b2) - else None)" -| "convert (TUInt b1) (TUInt b2) v = - (if b1 \ b2 - then Some (v, TUInt b2) - else None)" -| "convert (TUInt b1) (TSInt b2) v = - (if b1 < b2 - then Some (v, TSInt b2) - else None)" -| "convert TBool TBool v = Some (v, TBool)" -| "convert TAddr TAddr v = Some (v, TAddr)" -| "convert _ _ _ = None" - -lemma convert_id[simp]: - "convert tp tp kv = Some (kv, tp)" - by (metis Types.exhaust convert.simps(1) convert.simps(2) convert.simps(4) convert.simps(5) order_refl) - -(*Covered informally*) -fun olift :: - "(int \ int \ int) \ Types \ Types \ Valuetype \ Valuetype \ (Valuetype * Types) option" -where - "olift op (TSInt b1) (TSInt b2) v1 v2 = - Some (createSInt (max b1 b2) (op \v1\ \v2\), TSInt (max b1 b2))" -| "olift op (TUInt b1) (TUInt b2) v1 v2 = - Some (createUInt (max b1 b2) (op \v1\ \v2\), TUInt (max b1 b2))" -| "olift op (TSInt b1) (TUInt b2) v1 v2 = - (if b2 < b1 - then Some (createSInt b1 (op \v1\ \v2\), TSInt b1) - else None)" -| "olift op (TUInt b1) (TSInt b2) v1 v2 = - (if b1 < b2 - then Some (createSInt b2 (op \v1\ \v2\), TSInt b2) - else None)" -| "olift _ _ _ _ _ = None" - -(*Covered*) -fun plift :: - "(int \ int \ bool) \ Types \ Types \ Valuetype \ Valuetype \ (Valuetype * Types) option" -where - "plift op (TSInt b1) (TSInt b2) v1 v2 = Some (createBool (op \v1\ \v2\), TBool)" -| "plift op (TUInt b1) (TUInt b2) v1 v2 = Some (createBool (op \v1\ \v2\), TBool)" -| "plift op (TSInt b1) (TUInt b2) v1 v2 = - (if b2 < b1 - then Some (createBool (op \v1\ \v2\), TBool) - else None)" -| "plift op (TUInt b1) (TSInt b2) v1 v2 = - (if b1 < b2 - then Some (createBool (op \v1\ \v2\), TBool) - else None)" -| "plift _ _ _ _ _ = None" - -(*Covered*) -definition add :: "Types \ Types \ Valuetype \ Valuetype \ (Valuetype * Types) option" -where - "add = olift (+)" - -(*Covered informally*) -definition sub :: "Types \ Types \ Valuetype \ Valuetype \ (Valuetype * Types) option" -where - "sub = olift (-)" - -(*Covered informally*) -definition equal :: "Types \ Types \ Valuetype \ Valuetype \ (Valuetype * Types) option" -where - "equal = plift (=)" - -(*Covered informally*) -definition less :: "Types \ Types \ Valuetype \ Valuetype \ (Valuetype * Types) option" -where - "less = plift (<)" - -declare less_def [solidity_symbex] - -(*Covered informally*) -definition leq :: "Types \ Types \ Valuetype \ Valuetype \ (Valuetype * Types) option" -where - "leq = plift (\)" - -(*Covered*) -fun vtand :: "Types \ Types \ Valuetype \ Valuetype \ (Valuetype * Types) option" -where - "vtand TBool TBool a b = - (if a = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True \ b = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True then Some (ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True, TBool) - else Some (ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l False, TBool))" -| "vtand _ _ _ _ = None" - -(*Covered informally*) -fun vtor :: "Types \ Types \ Valuetype \ Valuetype \ (Valuetype * Types) option" -where - "vtor TBool TBool a b = - (if a = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l False \ b = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l False - then Some (ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l False, TBool) - else Some (ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True, TBool))" -| "vtor _ _ _ _ = None" - -(*Covered informally*) -fun ival :: "Types \ Valuetype" -where - "ival (TSInt x) = ShowL\<^sub>i\<^sub>n\<^sub>t 0" -| "ival (TUInt x) = ShowL\<^sub>i\<^sub>n\<^sub>t 0" -| "ival TBool = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l False" -| "ival TAddr = STR ''0x0000000000000000000000000000000000000000''" - -end +section\Value Types\ + +theory Valuetypes +imports ReadShow +begin + +fun iter :: "(int \ 'b \ 'b) \ 'b \ int \ 'b" +where + "iter f v x = (if x \ 0 then v + else f (x-1) (iter f v (x-1)))" + +fun iter' :: "(int \ 'b \ 'b option) \ 'b \ int \ 'b option" +where + "iter' f v x = (if x \ 0 then Some v + else case iter' f v (x-1) of + Some v' \ f (x-1) v' + | None \ None)" + +type_synonym Address = String.literal +type_synonym Location = String.literal +type_synonym Valuetype = String.literal + +(*Covered*) +datatype Types = TSInt nat + | TUInt nat + | TBool + | TAddr + +(*Covered*) +definition createSInt :: "nat \ int \ Valuetype" +where + "createSInt b v = + (if v \ 0 + then ShowL\<^sub>i\<^sub>n\<^sub>t (-(2^(b-1)) + (v+2^(b-1)) mod (2^b)) + else ShowL\<^sub>i\<^sub>n\<^sub>t (2^(b-1) - (-v+2^(b-1)-1) mod (2^b) - 1))" + +lemma upper_bound: + fixes b::nat + and c::int + assumes "b > 0" + and "c < 2^(b-1)" + shows "c + 2^(b-1) < 2^b" +proof - + have a1: "\P. (\b::nat. P b) \ (\b>0. P ((b-1)::nat))" by simp + have b2: "\b::nat. (\(c::int)<2^b. (c + 2^b) < 2^(Suc b))" by simp + show ?thesis using a1[OF b2] assms by simp +qed + +lemma upper_bound2: + fixes b::nat + and c::int + assumes "b > 0" + and "c < 2^b" + and "c \ 0" + shows "c - (2^(b-1)) < 2^(b-1)" +proof - + have a1: "\P. (\b::nat. P b) \ (\b>0. P ((b-1)::nat))" by simp + have b2: "\b::nat. (\(c::int)<2^(Suc b). c\0 \ (c - 2^b) < 2^b)" by simp + show ?thesis using a1[OF b2] assms by simp +qed + +lemma upper_bound3: + fixes b::nat + and v::int + defines "x \ - (2 ^ (b - 1)) + (v + 2 ^ (b - 1)) mod 2 ^ b" + assumes "b>0" + shows "x < 2^(b-1)" + using upper_bound2 assms by auto + +lemma lower_bound: + fixes b::nat + assumes "b>0" + shows "\(c::int) \ -(2^(b-1)). (-c + 2^(b-1) - 1 < 2^b)" +proof - + have a1: "\P. (\b::nat. P b) \ (\b>0. P ((b-1)::nat))" by simp + have b2: "\b::nat. \(c::int) \ -(2^b). (-c + (2^b) - 1) < 2^(Suc b)" by simp + show ?thesis using a1[OF b2] assms by simp +qed + +lemma lower_bound2: + fixes b::nat + and v::int + defines "x \ 2^(b - 1) - (-v+2^(b-1)-1) mod 2^b - 1" + assumes "b>0" + shows "x \ - (2 ^ (b - 1))" + using upper_bound2 assms by auto + +lemma createSInt_id_g0: + fixes b::nat + and v::int + assumes "v \ 0" + and "v < 2^(b-1)" + and "b > 0" + shows "createSInt b v = ShowL\<^sub>i\<^sub>n\<^sub>t v" +proof - + from assms have "v + 2^(b-1) \ 0" by simp + moreover from assms have "v + (2^(b-1)) < 2^b" using upper_bound[of b] by auto + ultimately have "(v + 2^(b-1)) mod (2^b) = v + 2^(b-1)" by simp + moreover from assms have "createSInt b v=ShowL\<^sub>i\<^sub>n\<^sub>t (-(2^(b-1)) + (v+2^(b-1)) mod (2^b))" unfolding createSInt_def by simp + ultimately show ?thesis by simp +qed + +lemma createSInt_id_l0: + fixes b::nat + and v::int + assumes "v < 0" + and "v \ -(2^(b-1))" + and "b > 0" + shows "createSInt b v = ShowL\<^sub>i\<^sub>n\<^sub>t v" +proof - + from assms have "-v + 2^(b-1) - 1 \ 0" by simp + moreover from assms have "-v + 2^(b-1) - 1 < 2^b" using lower_bound[of b] by auto + ultimately have "(-v + 2^(b-1) - 1) mod (2^b) = (-v + 2^(b-1) - 1)" by simp + moreover from assms have "createSInt b v= ShowL\<^sub>i\<^sub>n\<^sub>t (2^(b-1) - (-v+2^(b-1)-1) mod (2^b) - 1)" unfolding createSInt_def by simp + ultimately show ?thesis by simp +qed + +lemma createSInt_id: + fixes b::nat + and v::int + assumes "v < 2^(b-1)" + and "v \ -(2^(b-1))" + and "b > 0" + shows "createSInt b v = ShowL\<^sub>i\<^sub>n\<^sub>t v" using createSInt_id_g0 createSInt_id_l0 assms unfolding createSInt_def by simp + +(*Covered*) +definition createUInt :: "nat \ int \ Valuetype" + where "createUInt b v = ShowL\<^sub>i\<^sub>n\<^sub>t (v mod (2^b))" + +lemma createUInt_id: + assumes "v \ 0" + and "v < 2^b" + shows "createUInt b v = ShowL\<^sub>i\<^sub>n\<^sub>t v" +unfolding createUInt_def by (simp add: assms(1) assms(2)) + +definition createBool :: "bool \ Valuetype" +where + "createBool b = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l b" + +definition createAddress :: "Address \ Valuetype" +where + "createAddress ad = ad" + +definition checkSInt :: "nat \ Valuetype \ bool" +where + "checkSInt b v = ((foldr (\) (map is_digit (String.explode v)) True) \(ReadL\<^sub>i\<^sub>n\<^sub>t v \ -(2^(b-1)) \ ReadL\<^sub>i\<^sub>n\<^sub>t v < 2^(b-1)))" + +definition checkUInt :: "nat \ Valuetype \ bool" +where + "checkUInt b v = ((foldr (\) (map is_digit (String.explode v)) True) \ (ReadL\<^sub>i\<^sub>n\<^sub>t v \ 0 \ ReadL\<^sub>i\<^sub>n\<^sub>t v < 2^b))" + +fun convert :: "Types \ Types \ Valuetype \ Valuetype option" +where + "convert (TSInt b1) (TSInt b2) v = + (if b1 \ b2 + then Some v + else None)" +| "convert (TUInt b1) (TUInt b2) v = + (if b1 \ b2 + then Some v + else None)" +| "convert (TUInt b1) (TSInt b2) v = + (if b1 < b2 + then Some v + else None)" +| "convert TBool TBool v = Some v" +| "convert TAddr TAddr v = Some v" +| "convert _ _ _ = None" + +lemma convert_id[simp]: + "convert tp tp kv = Some kv" + by (metis Types.exhaust convert.simps(1) convert.simps(2) convert.simps(4) convert.simps(5) order_refl) + +(*Covered informally*) +fun olift :: + "(int \ int \ int) \ Types \ Types \ Valuetype \ Valuetype \ (Valuetype * Types) option" +where + "olift op (TSInt b1) (TSInt b2) v1 v2 = + Some (createSInt (max b1 b2) (op \v1\ \v2\), TSInt (max b1 b2))" +| "olift op (TUInt b1) (TUInt b2) v1 v2 = + Some (createUInt (max b1 b2) (op \v1\ \v2\), TUInt (max b1 b2))" +| "olift op (TSInt b1) (TUInt b2) v1 v2 = + (if b2 < b1 + then Some (createSInt b1 (op \v1\ \v2\), TSInt b1) + else None)" +| "olift op (TUInt b1) (TSInt b2) v1 v2 = + (if b1 < b2 + then Some (createSInt b2 (op \v1\ \v2\), TSInt b2) + else None)" +| "olift _ _ _ _ _ = None" + +(*Covered*) +fun plift :: + "(int \ int \ bool) \ Types \ Types \ Valuetype \ Valuetype \ (Valuetype * Types) option" +where + "plift op (TSInt b1) (TSInt b2) v1 v2 = Some (createBool (op \v1\ \v2\), TBool)" +| "plift op (TUInt b1) (TUInt b2) v1 v2 = Some (createBool (op \v1\ \v2\), TBool)" +| "plift op (TSInt b1) (TUInt b2) v1 v2 = + (if b2 < b1 + then Some (createBool (op \v1\ \v2\), TBool) + else None)" +| "plift op (TUInt b1) (TSInt b2) v1 v2 = + (if b1 < b2 + then Some (createBool (op \v1\ \v2\), TBool) + else None)" +| "plift _ _ _ _ _ = None" + +(*Covered*) +definition add :: "Types \ Types \ Valuetype \ Valuetype \ (Valuetype * Types) option" +where + "add = olift (+)" + +(*Covered informally*) +definition sub :: "Types \ Types \ Valuetype \ Valuetype \ (Valuetype * Types) option" +where + "sub = olift (-)" + +(*Covered informally*) +definition equal :: "Types \ Types \ Valuetype \ Valuetype \ (Valuetype * Types) option" +where + "equal = plift (=)" + +(*Covered informally*) +definition less :: "Types \ Types \ Valuetype \ Valuetype \ (Valuetype * Types) option" +where + "less = plift (<)" + +declare less_def [solidity_symbex] + +(*Covered informally*) +definition leq :: "Types \ Types \ Valuetype \ Valuetype \ (Valuetype * Types) option" +where + "leq = plift (\)" + +(*Covered*) +fun vtand :: "Types \ Types \ Valuetype \ Valuetype \ (Valuetype * Types) option" +where + "vtand TBool TBool a b = + (if a = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True \ b = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True then Some (ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True, TBool) + else Some (ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l False, TBool))" +| "vtand _ _ _ _ = None" + +(*Covered informally*) +fun vtor :: "Types \ Types \ Valuetype \ Valuetype \ (Valuetype * Types) option" +where + "vtor TBool TBool a b = + (if a = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l False \ b = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l False + then Some (ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l False, TBool) + else Some (ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True, TBool))" +| "vtor _ _ _ _ = None" + +definition checkBool :: "Valuetype \ bool" +where + "checkBool v = (if (v = STR ''True'' \ v = STR ''False'') then True else False)" + +definition checkAddress :: "Valuetype \ bool" + where + "checkAddress v = (if (size v = 42 \ ((String.explode v !1) = CHR ''x'')) then True else False)" + +(*value "checkBool STR ''True''"*) +(*value "checkAddress STR ''0x0000000000000000000000000000000000000000''"*) + +(*Covered informally*) +primrec ival :: "Types \ Valuetype" +where + "ival (TSInt x) = ShowL\<^sub>i\<^sub>n\<^sub>t 0" +| "ival (TUInt x) = ShowL\<^sub>i\<^sub>n\<^sub>t 0" +| "ival TBool = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l False" +| "ival TAddr = STR ''0x0000000000000000000000000000000000000000''" + + +declare convert.simps [simp del] +declare olift.simps [simp del] +declare plift.simps [simp del] +declare vtand.simps [simp del] +declare vtor.simps [simp del] + +end diff --git a/thys/Solidity/Weakest_Precondition.thy b/thys/Solidity/Weakest_Precondition.thy new file mode 100644 --- /dev/null +++ b/thys/Solidity/Weakest_Precondition.thy @@ -0,0 +1,2200 @@ +theory Weakest_Precondition + imports Solidity_Main +begin + +section \Setup for Monad VCG\ + +lemma wpstackvalue[wprule]: + assumes "\v. a = KValue v \ wp (f v) P E s" + and "\p. a = KCDptr p \ wp (g p) P E s" + and "\p. a = KMemptr p \ wp (h p) P E s" + and "\p. a = KStoptr p \ wp (i p) P E s" + shows "wp (case a of KValue v \ f v | KCDptr p \ g p | KMemptr p \ h p | KStoptr p \ i p) P E s" +using assms by (simp add: Stackvalue.split[of "\x. wp x P E s"]) + +lemma wpmtypes[wprule]: + assumes "\i m. a = MTArray i m \ wp (f i m) P E s" + and "\t. a = MTValue t \ wp (g t) P E s" + shows "wp (case a of MTArray i m \ f i m | MTValue t \ g t) P E s" +using assms by (simp add: MTypes.split[of "\x. wp x P E s"]) + +lemma wpstypes[wprule]: + assumes "\i m. a = STArray i m \ wp (f i m) P E s" + and "\t t'. a = STMap t t' \ wp (g t t') P E s" + and "\t. a = STValue t \ wp (h t) P E s" +shows "wp (case a of STArray i m \ f i m | STMap t t' \ g t t' | STValue t \ h t) P E s" +using assms by (simp add: STypes.split[of "\x. wp x P E s"]) + +lemma wptype[wprule]: + assumes "\v. a = Value v \ wp (f v) P E s" + and "\m. a = Calldata m \ wp (g m) P E s" + and "\m. a = Memory m \ wp (h m) P E s" + and "\t. a = Storage t \ wp (i t) P E s" +shows "wp (case a of Value v \ f v | Calldata m \ g m | Memory m \ h m | Storage s \ i s) P E s" + using assms by (simp add: Type.split[of "\x. wp x P E s"]) + +lemma wptypes[wprule]: +assumes "\x. a= TSInt x \ wp (f x) P E s" + and "\x. a = TUInt x \ wp (g x) P E s" + and "a = TBool \ wp h P E s" + and "a = TAddr \ wp i P E s" + shows "wp (case a of TSInt x \ f x | TUInt x \ g x | TBool \ h | TAddr \ i) P E s" +using assms by (simp add: Types.split[of "\x. wp x P E s"]) + +lemma wpltype[wprule]: + assumes "\l. a=LStackloc l \ wp (f l) P E s" + and "\l. a = LMemloc l \ wp (g l) P E s" + and "\l. a = LStoreloc l \ wp (h l) P E s" + shows "wp (case a of LStackloc l \ f l | LMemloc l \ g l | LStoreloc l \ h l) P E s" +using assms by (simp add: LType.split[of "\x. wp x P E s"]) + +lemma wpdenvalue[wprule]: + assumes "\l. a=Stackloc l \ wp (f l) P E s" + and "\l. a=Storeloc l \ wp (g l) P E s" + shows "wp (case a of Stackloc l \ f l | Storeloc l \ g l) P E s" + using assms by (simp add:Denvalue.split[of "\x. wp x P E s" f g a]) + +section \Calculus\ + +subsection \Hoare Triples\ + +type_synonym State_Predicate = "Accounts \ Stack \ MemoryT \ (Address \ StorageT) \ bool" + +definition validS :: "State_Predicate \ (unit, Ex ,State) state_monad \ State_Predicate \ (Ex \ bool) \ bool" + ("\_\\<^sub>S/ _ /(\_\\<^sub>S,/ \_\\<^sub>S)") +where + "\P\\<^sub>S f \Q\\<^sub>S,\E\\<^sub>S \ + \st. P (accounts st, stack st, memory st, storage st) + \ (case f st of + Normal (_,st') \ gas st' \ gas st \ Q (accounts st', stack st', memory st', storage st') + | Exception e \ E e)" + +definition wpS :: "(unit, Ex ,State) state_monad \ (State \ bool) \ (Ex \ bool) \ State \ bool" + where "wpS f P E st \ wp f (\_ st'. gas st' \ gas st \ P st') E st" + +lemma wpS_valid: + assumes "\st::State. P (accounts st, stack st, memory st, storage st) \ wpS f (\st. Q (accounts st, stack st, memory st, storage st)) E st" + shows "\P\\<^sub>S f \Q\\<^sub>S,\E\\<^sub>S" + unfolding validS_def + using assms unfolding wpS_def wp_def by simp + +lemma valid_wpS: + assumes "\P\\<^sub>S f \Q\\<^sub>S,\E\\<^sub>S" + shows "\st. P (accounts st, stack st, memory st, storage st) \ wpS f (\st. Q (accounts st, stack st, memory st, storage st))E st" + unfolding wpS_def wp_def using assms unfolding validS_def by simp + +context statement_with_gas +begin + +subsection \Skip\ + +lemma wp_Skip: + assumes "P (st\gas := gas st - costs SKIP ev cd st\)" + and "E Gas" + shows "wpS (\s. stmt SKIP ev cd s) P E st" + apply (subst stmt.simps(1)) + unfolding wpS_def + apply wp + using assms by auto + +subsection \Assign\ + +lemma wp_Assign: + fixes ex ev cd st lv + defines "ngas \ gas st - costs (ASSIGN lv ex) ev cd st" + assumes "\v t g l t' g' v'. + \expr ex ev cd (st\gas := ngas\) ngas = Normal ((KValue v, Value t), g); + lexp lv ev cd (st\gas := g\) g = Normal ((LStackloc l, Value t'), g'); + g' \ gas st; + convert t t' v = Some v'\ + \ P (st\gas := g', stack:=updateStore l (KValue v') (stack st)\)" + and "\v t g l t' g' v'. + \expr ex ev cd (st\gas := ngas\) ngas = Normal ((KValue v, Value t), g); + lexp lv ev cd (st\gas := g\) g = Normal ((LStoreloc l, Storage (STValue t')), g'); + g' \ gas st; + convert t t' v = Some v'\ + \ P (st\gas := g', storage:=(storage st) (address ev := (fmupd l v' (storage st (address ev))))\)" + and "\v t g l t' g' v' vt. + \expr ex ev cd (st\gas := ngas\) ngas = Normal ((KValue v, Value t), g); + lexp lv ev cd (st\gas := g\) g = Normal ((LMemloc l, Memory (MTValue t')), g'); + g' \ gas st; + convert t t' v = Some v'\ + \ P (st\gas := g', memory:=updateStore l (MValue v') (memory st)\)" + and "\p x t g l t' g' p' m'. + \expr ex ev cd (st\gas := ngas\) ngas = Normal ((KCDptr p, Calldata (MTArray x t)), g); + lexp lv ev cd (st\gas := g\) g = Normal ((LStackloc l, Memory t'), g'); + g' \ gas st; + accessStore l (stack st) = Some (KMemptr p'); + cpm2m p p' x t cd (memory st) = Some m'\ + \ P (st\gas := g', memory:=m'\)" + and "\p x t g l t' g' p' s'. + \expr ex ev cd (st\gas := ngas\) ngas = Normal ((KCDptr p, Calldata (MTArray x t)), g); + lexp lv ev cd (st\gas := g\) g = Normal ((LStackloc l, Storage t'), g'); + g' \ gas st; + accessStore l (stack st) = Some (KStoptr p'); + cpm2s p p' x t cd (storage st (address ev)) = Some s'\ + \ P (st\gas := g', storage:=(storage st) (address ev := s')\)" + and "\p x t g l t' g' s'. + \expr ex ev cd (st\gas := ngas\) ngas = Normal ((KCDptr p, Calldata (MTArray x t)), g); + lexp lv ev cd (st\gas := g\) g = Normal ((LStoreloc l, t'), g'); + g' \ gas st; + cpm2s p l x t cd (storage st (address ev)) = Some s'\ + \ P (st\gas := g', storage:=(storage st) (address ev := s')\)" + and "\p x t g l t' g' m'. + \expr ex ev cd (st\gas := ngas\) ngas = Normal ((KCDptr p, Calldata (MTArray x t)), g); + lexp lv ev cd (st\gas := g\) g = Normal ((LMemloc l, t'), g'); + g' \ gas st; + cpm2m p l x t cd (memory st) = Some m'\ + \ P (st\gas := g', memory:=m'\)" + and "\p x t g l t' g'. + \expr ex ev cd (st\gas := ngas\) ngas = Normal ((KMemptr p, Memory (MTArray x t)), g); + lexp lv ev cd (st\gas := g\) g = Normal ((LStackloc l, Memory t'), g'); + g' \ gas st\ + \ P (st\gas := g', stack:=updateStore l (KMemptr p) (stack st)\)" + and "\p x t g l t' g' p' s'. + \expr ex ev cd (st\gas := ngas\) ngas = Normal ((KMemptr p, Memory (MTArray x t)), g); + lexp lv ev cd (st\gas := g\) g = Normal ((LStackloc l, Storage t'), g'); + g' \ gas st; + accessStore l (stack st) = Some (KStoptr p'); + cpm2s p p' x t (memory st) (storage st (address ev)) = Some s'\ + \ P (st\gas := g', storage:=(storage st) (address ev := s')\)" + and "\p x t g l t' g' s'. + \expr ex ev cd (st\gas := ngas\) ngas = Normal ((KMemptr p, Memory (MTArray x t)), g); + lexp lv ev cd (st\gas := g\) g = Normal ((LStoreloc l, t'), g'); + g' \ gas st; + cpm2s p l x t (memory st) (storage st (address ev)) = Some s'\ + \ P (st\gas := g', storage:=(storage st) (address ev := s')\)" + and "\p x t g l t' g'. + \expr ex ev cd (st\gas := ngas\) ngas = Normal ((KMemptr p, Memory (MTArray x t)), g); + lexp lv ev cd (st\gas := g\) g = Normal ((LMemloc l, t'), g'); + g' \ gas st\ + \ P (st\gas := g', memory:=updateStore l (MPointer p) (memory st)\)" + and "\p x t g l t' g' p' m'. + \expr ex ev cd (st\gas := ngas\) ngas = Normal ((KStoptr p, Storage (STArray x t)), g); + lexp lv ev cd (st\gas := g\) g = Normal ((LStackloc l, Memory t'), g'); + g' \ gas st; + accessStore l (stack st) = Some (KMemptr p'); + cps2m p p' x t (storage st (address ev)) (memory st) = Some m'\ + \ P (st\gas := g', memory:=m'\)" + and "\p x t g l t' g'. + \expr ex ev cd (st\gas := ngas\) ngas = Normal ((KStoptr p, Storage (STArray x t)), g); + lexp lv ev cd (st\gas := g\) g = Normal ((LStackloc l, Storage t'), g'); + g' \ gas st\ + \ P (st\gas := g', stack:=updateStore l (KStoptr p) (stack st)\)" + and "\p x t g l t' g' s'. + \expr ex ev cd (st\gas := ngas\) ngas = Normal ((KStoptr p, Storage (STArray x t)), g); + lexp lv ev cd (st\gas := g\) g = Normal ((LStoreloc l, t'), g'); + g' \ gas st; + copy p l x t (storage st (address ev)) = Some s'\ + \ P (st\gas := g', storage:=(storage st) (address ev := s')\)" + and "\p x t g l t' g' m'. + \expr ex ev cd (st\gas := ngas\) ngas = Normal ((KStoptr p, Storage (STArray x t)), g); + lexp lv ev cd (st\gas := g\) g = Normal ((LMemloc l, t'), g'); + g' \ gas st; + cps2m p l x t (storage st (address ev)) (memory st) = Some m'\ + \ P (st\gas := g', memory:=m'\)" + and "\p t t' g l t'' g'. + \expr ex ev cd (st\gas := ngas\) ngas = Normal ((KStoptr p, Storage (STMap t t')), g); + lexp lv ev cd (st\gas := g\) g = Normal ((LStackloc l, t''), g'); + g' \ gas st\ + \ P (st\gas := g', stack:=updateStore l (KStoptr p) (stack st)\)" + and "E Gas" + and "E Err" + shows "wpS (\s. stmt (ASSIGN lv ex) ev cd s) P E st" + apply (subst stmt.simps(2)) + unfolding wpS_def + apply wp + apply (simp_all add: Ex.induct[OF assms(18,19)]) +proof - + fix a g aa b v t ab ga ac ba l t' v' + assume "costs (ASSIGN lv ex) ev cd st < gas st" + and *: "local.expr ex ev cd (st\gas := gas st - costs (ASSIGN lv ex) ev cd st\) (gas st - costs (ASSIGN lv ex) ev cd st) = Normal ((KValue v, Value t), g)" + and "a = (KValue v, Value t)" + and "aa = KValue v" + and "b = Value t" + and **: "local.lexp lv ev cd (st\gas := g\) g = Normal ((LStackloc l, Value t'), ga)" + and "ab = (LStackloc l, Value t')" + and "ac = LStackloc l" + and "ba = Value t'" + and "convert t t' v = Some v'" + moreover have "ga \ gas st" + proof - + have "ga \ g" using lexp_gas[OF **] by simp + also have "\ \ gas st" using msel_ssel_expr_load_rexp_gas(3)[OF *] by simp + finally show ?thesis . + qed + ultimately show "ga \ gas st \ P (st\gas := ga, stack := updateStore l (KValue v') (stack st)\)" using assms(2) unfolding ngas_def by simp +next + fix a g aa b v t ab ga ac ba l MTypes t' v' + assume "costs (ASSIGN lv ex) ev cd st < gas st" + and *: "local.expr ex ev cd (st\gas := gas st - costs (ASSIGN lv ex) ev cd st\) (gas st - costs (ASSIGN lv ex) ev cd st) = Normal ((KValue v, Value t), g)" + and "a = (KValue v, Value t)" + and "aa = KValue v" + and "b = Value t" + and **: "local.lexp lv ev cd (st\gas := g\) g = Normal ((LMemloc l, Memory (MTValue t')), ga)" + and "ab = (LMemloc l, Memory (MTValue t'))" + and "ac = LMemloc l" + and "ba = Memory (MTValue t')" + and "MTypes = MTValue t'" + and "convert t t' v = Some v'" + moreover have "ga \ gas st" + proof - + have "ga \ g" using lexp_gas[OF **] by simp + also have "\ \ gas st" using msel_ssel_expr_load_rexp_gas(3)[OF *] by simp + finally show ?thesis . + qed + ultimately show "ga \ gas st \ P (st\gas := ga, memory := updateStore l (MValue v') (memory st)\)" using assms(4) unfolding ngas_def by simp +next + fix a g aa b v t ab ga ac ba l ta Types v' + assume "costs (ASSIGN lv ex) ev cd st < gas st" + and *: "local.expr ex ev cd (st\gas := gas st - costs (ASSIGN lv ex) ev cd st\) (gas st - costs (ASSIGN lv ex) ev cd st) = Normal ((KValue v, Value t), g)" + and "a = (KValue v, Value t)" + and "aa = KValue v" + and "b = Value t" + and **: "local.lexp lv ev cd (st\gas := g\) g = Normal ((LStoreloc l, Storage (STValue Types)), ga)" + and "ab = (LStoreloc l, Storage (STValue Types))" + and "ac = LStoreloc l" + and "ba = Storage (STValue Types)" + and "ta = STValue Types" + and "convert t Types v = Some v'" + moreover have "ga \ gas st" + proof - + have "ga \ g" using lexp_gas[OF **] by simp + also have "\ \ gas st" using msel_ssel_expr_load_rexp_gas(3)[OF *] by simp + finally show ?thesis . + qed + ultimately show "ga \ gas st \ P (st\gas := ga, storage := (storage st) (address ev := fmupd l v' (storage st (address ev)))\)" using assms(3) unfolding ngas_def by simp +next + fix a g aa b p MTypes x t ab ga ac xa l MTypesa y literal ya + assume "costs (ASSIGN lv ex) ev cd st < gas st" + and *: "local.expr ex ev cd (st\gas := gas st - costs (ASSIGN lv ex) ev cd st\) (gas st - costs (ASSIGN lv ex) ev cd st) = Normal ((KCDptr p, Calldata (MTArray x t)), g)" + and "a = (KCDptr p, Calldata (MTArray x t))" + and "aa = KCDptr p" + and "b = Calldata (MTArray x t)" + and "MTypes = MTArray x t" + and **: "local.lexp lv ev cd (st\gas := g\) g = Normal ((LStackloc l, Memory MTypesa), ga)" + and "ab = (LStackloc l, Memory MTypesa)" + and "ac = LStackloc l" + and "xa = Memory MTypesa" + and "accessStore l (stack st) = Some (KMemptr literal)" + and "y = KMemptr literal" + and "cpm2m p literal x t cd (memory st) = Some ya" + moreover have "ga \ gas st" + proof - + have "ga \ g" using lexp_gas[OF **] by simp + also have "\ \ gas st" using msel_ssel_expr_load_rexp_gas(3)[OF *] by simp + finally show ?thesis . + qed + ultimately show "ga \ gas st \ P (st\gas := ga, memory := ya\)" using assms(5) unfolding ngas_def by simp +next + fix a g aa b p MTypes x t ab ga ac xa l ta y literal ya + assume "costs (ASSIGN lv ex) ev cd st < gas st" + and *: "local.expr ex ev cd (st\gas := gas st - costs (ASSIGN lv ex) ev cd st\) (gas st - costs (ASSIGN lv ex) ev cd st) = Normal ((KCDptr p, Calldata (MTArray x t)), g)" + and "a = (KCDptr p, Calldata (MTArray x t))" + and "aa = KCDptr p" + and "b = Calldata (MTArray x t)" + and "MTypes = MTArray x t" + and **: "local.lexp lv ev cd (st\gas := g\) g = Normal ((LStackloc l, Storage ta), ga)" + and "ab = (LStackloc l, Storage ta)" + and "ac = LStackloc l" + and "xa = Storage ta" + and "accessStore l (stack st) = Some (KStoptr literal)" + and "y = KStoptr literal" + and "cpm2s p literal x t cd (storage st (address ev)) = Some ya" + moreover have "ga \ gas st" + proof - + have "ga \ g" using lexp_gas[OF **] by simp + also have "\ \ gas st" using msel_ssel_expr_load_rexp_gas(3)[OF *] by simp + finally show ?thesis . + qed + ultimately show "ga \ gas st \ P (st\gas := ga, storage := (storage st) (address ev := ya)\)" using assms(6) unfolding ngas_def by simp +next + fix a g aa b p MTypes x t ab ga ac xa l y + assume "costs (ASSIGN lv ex) ev cd st < gas st" + and *: "local.expr ex ev cd (st\gas := gas st - costs (ASSIGN lv ex) ev cd st\) (gas st - costs (ASSIGN lv ex) ev cd st) = Normal ((KCDptr p, Calldata (MTArray x t)), g)" + and "a = (KCDptr p, Calldata (MTArray x t))" + and "aa = KCDptr p" + and "b = Calldata (MTArray x t)" + and "MTypes = MTArray x t" + and **: "local.lexp lv ev cd (st\gas := g\) g = Normal ((LMemloc l, xa), ga)" + and "ab = (LMemloc l, xa)" + and "ac = LMemloc l" + and "cpm2m p l x t cd (memory st) = Some y" + moreover have "ga \ gas st" + proof - + have "ga \ g" using lexp_gas[OF **] by simp + also have "\ \ gas st" using msel_ssel_expr_load_rexp_gas(3)[OF *] by simp + finally show ?thesis . + qed + ultimately show "ga \ gas st \ P (st\gas := ga, memory := y\)" using assms(8) unfolding ngas_def by simp +next + fix a g aa b p MTypes x t ab ga ac xa l y + assume "costs (ASSIGN lv ex) ev cd st < gas st" + and *: "local.expr ex ev cd (st\gas := gas st - costs (ASSIGN lv ex) ev cd st\) (gas st - costs (ASSIGN lv ex) ev cd st) = Normal ((KCDptr p, Calldata (MTArray x t)), g)" + and "a = (KCDptr p, Calldata (MTArray x t))" + and "aa = KCDptr p" + and "b = Calldata (MTArray x t)" + and "MTypes = MTArray x t" + and **: "local.lexp lv ev cd (st\gas := g\) g = Normal ((LStoreloc l, xa), ga)" + and "ab = (LStoreloc l, xa)" + and "ac = LStoreloc l" + and "cpm2s p l x t cd (storage st (address ev)) = Some y" + moreover have "ga \ gas st" + proof - + have "ga \ g" using lexp_gas[OF **] by simp + also have "\ \ gas st" using msel_ssel_expr_load_rexp_gas(3)[OF *] by simp + finally show ?thesis . + qed + ultimately show "ga \ gas st \ P (st\gas := ga, storage := (storage st) (address ev := y)\)" using assms(7) unfolding ngas_def by simp +next + fix a g aa b p MTypes x t ab ga ac xa l MTypesa + assume "costs (ASSIGN lv ex) ev cd st < gas st" + and *: "local.expr ex ev cd (st\gas := gas st - costs (ASSIGN lv ex) ev cd st\) (gas st - costs (ASSIGN lv ex) ev cd st) = Normal ((KMemptr p, Memory (MTArray x t)), g)" + and "a = (KMemptr p, Memory (MTArray x t))" + and "aa = KMemptr p" + and "b = Memory (MTArray x t)" + and "MTypes = MTArray x t" + and **: "local.lexp lv ev cd (st\gas := g\) g = Normal ((LStackloc l, Memory MTypesa), ga)" + and "ab = (LStackloc l, Memory MTypesa)" + and "ac = LStackloc l" + and "xa = Memory MTypesa" + moreover have "ga \ gas st" + proof - + have "ga \ g" using lexp_gas[OF **] by simp + also have "\ \ gas st" using msel_ssel_expr_load_rexp_gas(3)[OF *] by simp + finally show ?thesis . + qed + ultimately show "ga \ gas st \ P (st\gas := ga, stack := updateStore l (KMemptr p) (stack st)\)" using assms(9) unfolding ngas_def by simp +next + fix a g aa b p MTypes x t ab ga ac xa l ta y literal ya + assume "costs (ASSIGN lv ex) ev cd st < gas st" + and *: "local.expr ex ev cd (st\gas := gas st - costs (ASSIGN lv ex) ev cd st\) (gas st - costs (ASSIGN lv ex) ev cd st) = Normal ((KMemptr p, Memory (MTArray x t)), g)" + and "a = (KMemptr p, Memory (MTArray x t))" + and "aa = KMemptr p" + and "b = Memory (MTArray x t)" + and "MTypes = MTArray x t" + and **: "local.lexp lv ev cd (st\gas := g\) g = Normal ((LStackloc l, Storage ta), ga)" + and "ab = (LStackloc l, Storage ta)" + and "ac = LStackloc l" + and "xa = Storage ta" + and "accessStore l (stack st) = Some (KStoptr literal)" + and "y = KStoptr literal" + and "cpm2s p literal x t (memory st) (storage st (address ev)) = Some ya" + moreover have "ga \ gas st" + proof - + have "ga \ g" using lexp_gas[OF **] by simp + also have "\ \ gas st" using msel_ssel_expr_load_rexp_gas(3)[OF *] by simp + finally show ?thesis . + qed + ultimately show "ga \ gas st \ P (st\gas := ga, storage := (storage st) (address ev := ya)\)" using assms(10) unfolding ngas_def by simp +next + fix a g aa b p MTypes x t ab ga ac xa l + assume "costs (ASSIGN lv ex) ev cd st < gas st" + and *: "local.expr ex ev cd (st\gas := gas st - costs (ASSIGN lv ex) ev cd st\) (gas st - costs (ASSIGN lv ex) ev cd st) = Normal ((KMemptr p, Memory (MTArray x t)), g)" + and "a = (KMemptr p, Memory (MTArray x t))" + and "aa = KMemptr p" + and "b = Memory (MTArray x t)" + and "MTypes = MTArray x t" + and **: "local.lexp lv ev cd (st\gas := g\) g = Normal ((LMemloc l, xa), ga)" + and "ab = (LMemloc l, xa)" + and "ac = LMemloc l" + moreover have "ga \ gas st" + proof - + have "ga \ g" using lexp_gas[OF **] by simp + also have "\ \ gas st" using msel_ssel_expr_load_rexp_gas(3)[OF *] by simp + finally show ?thesis . + qed + ultimately show "ga \ gas st \ P (st\gas := ga, memory := updateStore l (MPointer p) (memory st)\)" using assms(12) unfolding ngas_def by simp +next + fix a g aa b p MTypes x t ab ga ac xa l y + assume "costs (ASSIGN lv ex) ev cd st < gas st" + and *: "local.expr ex ev cd (st\gas := gas st - costs (ASSIGN lv ex) ev cd st\) (gas st - costs (ASSIGN lv ex) ev cd st) = Normal ((KMemptr p, Memory (MTArray x t)), g)" + and "a = (KMemptr p, Memory (MTArray x t))" + and "aa = KMemptr p" + and "b = Memory (MTArray x t)" + and "MTypes = MTArray x t" + and **: "local.lexp lv ev cd (st\gas := g\) g = Normal ((LStoreloc l, xa), ga)" + and "ab = (LStoreloc l, xa)" + and "ac = LStoreloc l" + and "cpm2s p l x t (memory st) (storage st (address ev)) = Some y" + moreover have "ga \ gas st" + proof - + have "ga \ g" using lexp_gas[OF **] by simp + also have "\ \ gas st" using msel_ssel_expr_load_rexp_gas(3)[OF *] by simp + finally show ?thesis . + qed + ultimately show "ga \ gas st \ P (st\gas := ga, storage := (storage st) (address ev := y)\)" using assms(11) unfolding ngas_def by simp +next + fix a g aa b p t x ta ab ga ac xa l MTypes y literal ya + assume "costs (ASSIGN lv ex) ev cd st < gas st" + and *: "local.expr ex ev cd (st\gas := gas st - costs (ASSIGN lv ex) ev cd st\) (gas st - costs (ASSIGN lv ex) ev cd st) = Normal ((KStoptr p, Storage (STArray x ta)), g)" + and "a = (KStoptr p, Storage (STArray x ta))" + and "aa = KStoptr p" + and "b = Storage (STArray x ta)" + and "t = STArray x ta" + and **: "local.lexp lv ev cd (st\gas := g\) g = Normal ((LStackloc l, Memory MTypes), ga)" + and "ab = (LStackloc l, Memory MTypes)" + and "ac = LStackloc l" + and "xa = Memory MTypes" + and "accessStore l (stack st) = Some (KMemptr literal)" + and "y = KMemptr literal" + and "cps2m p literal x ta (storage st (address ev)) (memory st) = Some ya" + moreover have "ga \ gas st" + proof - + have "ga \ g" using lexp_gas[OF **] by simp + also have "\ \ gas st" using msel_ssel_expr_load_rexp_gas(3)[OF *] by simp + finally show ?thesis . + qed + ultimately show "ga \ gas st \ P (st\gas := ga, memory := ya\)" using assms(13) unfolding ngas_def by simp +next + fix a g aa b p t x ta ab ga ac xa l tb + assume "costs (ASSIGN lv ex) ev cd st < gas st" + and *: "local.expr ex ev cd (st\gas := gas st - costs (ASSIGN lv ex) ev cd st\) (gas st - costs (ASSIGN lv ex) ev cd st) = Normal ((KStoptr p, Storage (STArray x ta)), g)" + and "a = (KStoptr p, Storage (STArray x ta))" + and "aa = KStoptr p" + and "b = Storage (STArray x ta)" + and "t = STArray x ta" + and **: "local.lexp lv ev cd (st\gas := g\) g = Normal ((LStackloc l, Storage tb), ga)" + and "ab = (LStackloc l, Storage tb)" + and "ac = LStackloc l" + and "xa = Storage tb" + moreover have "ga \ gas st" + proof - + have "ga \ g" using lexp_gas[OF **] by simp + also have "\ \ gas st" using msel_ssel_expr_load_rexp_gas(3)[OF *] by simp + finally show ?thesis . + qed + ultimately show "ga \ gas st \ P (st\gas := ga, stack := updateStore l (KStoptr p) (stack st)\)" using assms(14) unfolding ngas_def by simp +next + fix a g aa b p t x ta ab ga ac xa l y + assume "costs (ASSIGN lv ex) ev cd st < gas st" + and *: "local.expr ex ev cd (st\gas := gas st - costs (ASSIGN lv ex) ev cd st\) (gas st - costs (ASSIGN lv ex) ev cd st) = Normal ((KStoptr p, Storage (STArray x ta)), g)" + and "a = (KStoptr p, Storage (STArray x ta))" + and "aa = KStoptr p" + and "b = Storage (STArray x ta)" + and "t = STArray x ta" + and **: "local.lexp lv ev cd (st\gas := g\) g = Normal ((LMemloc l, xa), ga)" + and "ab = (LMemloc l, xa)" + and "ac = LMemloc l" + and "cps2m p l x ta (storage st (address ev)) (memory st) = Some y" + moreover have "ga \ gas st" + proof - + have "ga \ g" using lexp_gas[OF **] by simp + also have "\ \ gas st" using msel_ssel_expr_load_rexp_gas(3)[OF *] by simp + finally show ?thesis . + qed + ultimately show "ga \ gas st \ P (st\gas := ga, memory := y\)" using assms(16) unfolding ngas_def by simp +next + fix a g aa b p t x ta ab ga ac xa l y + assume "costs (ASSIGN lv ex) ev cd st < gas st" + and *: "local.expr ex ev cd (st\gas := gas st - costs (ASSIGN lv ex) ev cd st\) (gas st - costs (ASSIGN lv ex) ev cd st) = Normal ((KStoptr p, Storage (STArray x ta)), g)" + and "a = (KStoptr p, Storage (STArray x ta))" + and "aa = KStoptr p" + and "b = Storage (STArray x ta)" + and "t = STArray x ta" + and **: "local.lexp lv ev cd (st\gas := g\) g = Normal ((LStoreloc l, xa), ga)" + and "ab = (LStoreloc l, xa)" + and "ac = LStoreloc l" + and "copy p l x ta (storage st (address ev)) = Some y" + moreover have "ga \ gas st" + proof - + have "ga \ g" using lexp_gas[OF **] by simp + also have "\ \ gas st" using msel_ssel_expr_load_rexp_gas(3)[OF *] by simp + finally show ?thesis . + qed + ultimately show "ga \ gas st \ P (st\gas := ga, storage := (storage st) (address ev := y)\)" using assms(15) unfolding ngas_def by simp +next + fix a g aa b p t ta t' ab ga ac x l + assume " costs (ASSIGN lv ex) ev cd st < gas st" + and *: "local.expr ex ev cd (st\gas := gas st - costs (ASSIGN lv ex) ev cd st\) (gas st - costs (ASSIGN lv ex) ev cd st) = Normal ((KStoptr p, Storage (STMap ta t')), g)" + and "a = (KStoptr p, Storage (STMap ta t'))" + and "aa = KStoptr p" + and "b = Storage (STMap ta t')" + and "t = STMap ta t'" + and **: "local.lexp lv ev cd (st\gas := g\) g = Normal ((LStackloc l, x), ga)" + and "ab = (LStackloc l, x)" + and "ac = LStackloc l" + moreover have "ga \ gas st" + proof - + have "ga \ g" using lexp_gas[OF **] by simp + also have "\ \ gas st" using msel_ssel_expr_load_rexp_gas(3)[OF *] by simp + finally show ?thesis . + qed + ultimately show "ga \ gas st \ P (st\gas := ga, stack := updateStore l (KStoptr p) (stack st)\)" using assms(17) unfolding ngas_def by simp +qed + +subsection \Composition\ + +lemma wp_Comp: + assumes "wpS (stmt s1 ev cd) (\st. wpS (stmt s2 ev cd) P E st) E (st\gas := gas st - costs (COMP s1 s2) ev cd st\)" + and "E Gas" + and "E Err" + shows "wpS (\s. stmt (COMP s1 s2) ev cd s) P E st" + apply (subst stmt.simps(3)) + unfolding wpS_def + apply wp + using assms unfolding wpS_def wp_def by (auto split:result.split) + +subsection \Conditional\ + +lemma wp_ITE: + assumes "\g g'. expr ex ev cd (st\gas := g\) g = Normal ((KValue \True\, Value TBool), g') \ wpS (stmt s1 ev cd) P E (st\gas := g'\)" + and "\g g'. expr ex ev cd (st\gas := g\) g = Normal ((KValue \False\, Value TBool), g') \ wpS (stmt s2 ev cd) P E (st\gas := g'\)" + and "E Gas" + and "E Err" + shows "wpS (\s. stmt (ITE ex s1 s2) ev cd s) P E st" + apply (subst stmt.simps(4)) + unfolding wpS_def + apply wp + apply (simp_all add: Ex.induct[OF assms(3,4)]) + proof - + fix a g aa ba b v + assume "costs (ITE ex s1 s2) ev cd st < gas st" + and *: "local.expr ex ev cd (st\gas := gas st - costs (ITE ex s1 s2) ev cd st\) (gas st - costs (ITE ex s1 s2) ev cd st) = Normal ((KValue \True\, Value TBool), g)" + and "a = (KValue \True\, Value TBool)" + and "aa = KValue \True\" and "ba = Value TBool" and "v = TBool" and "b = \True\" + then have "wpS (stmt s1 ev cd) P E (st\gas := g\)" using assms(1) by simp + moreover have "g \ gas st" using msel_ssel_expr_load_rexp_gas(3)[OF *] by simp + ultimately show "wp (local.stmt s1 ev cd) (\_ st'. gas st' \ gas st \ P st') E (st\gas := g\)" + unfolding wpS_def wp_def by (simp split:result.split_asm prod.split_asm) + next + fix a g aa ba b v + assume "costs (ITE ex s1 s2) ev cd st < gas st" + and *: "local.expr ex ev cd (st\gas := gas st - costs (ITE ex s1 s2) ev cd st\) (gas st - costs (ITE ex s1 s2) ev cd st) = Normal ((KValue \False\, Value TBool), g)" + and "a = (KValue \False\, Value TBool)" + and "aa = KValue \False\" and "ba = Value TBool" and "v = TBool" and "b = \False\" + then have "wpS (stmt s2 ev cd) P E (st\gas := g\)" using assms(2) by simp + moreover have "g \ gas st" using msel_ssel_expr_load_rexp_gas(3)[OF *] by simp + ultimately show "wp (local.stmt s2 ev cd) (\_ st'. gas st' \ gas st \ P st') E (st\gas := g\)" + unfolding wpS_def wp_def by (simp split:result.split_asm prod.split_asm) + qed + +subsection \While Loop\ + +lemma wp_While[rule_format]: + fixes iv::"Accounts \ Stack \ MemoryT \ (Address \ StorageT) \ bool" + assumes "\a k m s st g. \iv (a, k, m, s); expr ex ev cd (st\gas := gas st - costs (WHILE ex sm) ev cd st\) (gas st - costs (WHILE ex sm) ev cd st) = Normal ((KValue \False\, Value TBool), g)\ \ P (st\gas := g\)" + and "\a k m s st g. \iv (a, k, m, s); expr ex ev cd (st\gas := gas st - costs (WHILE ex sm) ev cd st\) (gas st - costs (WHILE ex sm) ev cd st) = Normal ((KValue \True\, Value TBool), g)\ \ wpS (stmt sm ev cd) (\st. iv (accounts st, stack st, memory st, storage st)) E (st\gas:=g\)" + and "E Err" + and "E Gas" + shows "iv (accounts st, stack st, memory st, storage st) \ wpS (\s. stmt (WHILE ex sm) ev cd s) P E st" +proof (induction st rule:gas_induct) + case (1 st) + show ?case + unfolding wpS_def wp_def + proof (split result.split, rule conjI; rule allI; rule impI) + fix x1 assume *: "local.stmt (WHILE ex sm) ev cd st = Normal x1" + then obtain b g where **: "expr ex ev cd (st\gas := gas st - costs (WHILE ex sm) ev cd st\) (gas st - costs (WHILE ex sm) ev cd st) = Normal ((KValue b, Value TBool), g)" by (simp only: stmt.simps, simp split:if_split_asm result.split_asm prod.split_asm Stackvalue.split_asm Type.split_asm Types.split_asm) + with * consider (t) "b = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True" | (f) "b = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l False" by (simp add:stmt.simps split:if_split_asm result.split_asm prod.split_asm Stackvalue.split_asm Type.split_asm Types.split_asm) + then show "iv (accounts st, stack st, memory st, storage st) \ (case x1 of (r, s') \ gas s' \ gas st \ P s')" + proof cases + case t + then obtain st' where ***: "stmt sm ev cd (st\gas := g\) = Normal ((), st')" using * ** by (auto simp add:stmt.simps split:if_split_asm result.split_asm) + then have ****: "local.stmt (WHILE ex sm) ev cd st' = Normal x1" using * ** t by (simp add:stmt.simps split:if_split_asm) + show ?thesis + proof + assume "iv (accounts st, stack st, memory st, storage st)" + then have "wpS (local.stmt sm ev cd) (\st. iv (accounts st, stack st, memory st, storage st)) E (st\gas:=g\)" using assms(2) ** t by auto + then have "iv (accounts st', stack st', memory st', storage st')" unfolding wpS_def wp_def using *** by (simp split:result.split_asm)+ + moreover have "gas st \ costs (WHILE ex sm) ev cd st" using * by (simp add:stmt.simps split:if_split_asm) + then have "gas st' < gas st" using stmt_gas[OF ***] msel_ssel_expr_load_rexp_gas(3)[OF **] while_not_zero[of ex sm ev cd st] by simp + ultimately have "wpS (local.stmt (WHILE ex sm) ev cd) P E st'" using 1 by simp + then show "(case x1 of (r, s') \ gas s' \ gas st \ P s')" unfolding wpS_def wp_def using **** `gas st' < gas st` by auto + qed + next + case f + show ?thesis + proof + assume "iv (accounts st, stack st, memory st, storage st)" + then have "P (st\gas := g\)" using ** assms(1) f by simp + moreover have "x1 = ((),st\gas := g\)" using * ** f by (simp add:stmt.simps true_neq_false[symmetric] split:if_split_asm) + moreover have "g \ gas st" using msel_ssel_expr_load_rexp_gas(3)[OF **] by simp + ultimately show "case x1 of (r, s') \ gas s' \ gas st \ P s'" by (auto split:prod.split) + qed + qed + next + fix x2 + show "iv (accounts st, stack st, memory st, storage st) \ E x2" using assms(3,4) Ex.nchotomy by metis + qed +qed + +subsection \Blocks\ + +lemma wp_blockNone: + assumes "\cd' mem' sck' e'. decl id0 tp None False cd (memory (st\gas := gas st - costs (BLOCK ((id0, tp), None) stm) ev cd st\)) (storage (st\gas := gas st - costs (BLOCK ((id0, tp), None) stm) ev cd st\)) + (cd, memory (st\gas := gas st - costs (BLOCK ((id0, tp), None) stm) ev cd st\), stack (st\gas := gas st - costs (BLOCK ((id0, tp), None) stm) ev cd st\), ev) = Some (cd', mem', sck', e') + \ wpS (stmt stm e' cd') P E (st\gas := gas st - costs (BLOCK ((id0, tp), None) stm) ev cd st, stack := sck', memory := mem'\)" + and "E Gas" + and "E Err" + shows "wpS (\s. stmt (BLOCK ((id0, tp), None) stm) ev cd s) P E st" + unfolding wpS_def wp_def +proof ((split result.split | split prod.split)+, rule conjI; (rule allI | rule impI)+) + fix x1 x1a x2 + assume "x1 = (x1a, x2)" + and "local.stmt (BLOCK ((id0, tp), None) stm) ev cd st = Normal x1" + then have "local.stmt (BLOCK ((id0, tp), None) stm) ev cd st = Normal (x1a, x2)" by simp + then show "gas x2 \ gas st \ P x2" + proof (cases rule: blockNone) + case (1 cd' mem' sck' e') + then show ?thesis using assms(1)[OF 1(2)] unfolding wpS_def wp_def by auto + qed +next + fix e + assume "local.stmt (BLOCK ((id0, tp), None) stm) ev cd st = Exception e" + show "E e" using assms(2,3) by (induction rule: Ex.induct) +qed + +lemma wp_blockSome: + assumes "\v t g' cd' mem' sck' e'. + \ expr ex' ev cd (st\gas := gas st - costs (BLOCK ((id0, tp), Some ex') stm) ev cd st\) (gas st - costs (BLOCK ((id0, tp), Some ex') stm) ev cd st) = Normal ((v, t), g'); + g' \ gas st - costs (BLOCK ((id0, tp), Some ex') stm) ev cd st; + decl id0 tp (Some (v,t)) False cd (memory st) (storage st) (cd, memory st, stack st, ev) = Some (cd', mem', sck', e')\ + \ wpS (stmt stm e' cd') P E (st\gas := g', stack := sck', memory := mem'\)" + and "E Gas" + and "E Err" + shows "wpS (\s. stmt (BLOCK ((id0, tp), Some ex') stm) ev cd s) P E st" + unfolding wpS_def wp_def +proof ((split result.split | split prod.split)+, rule conjI; (rule allI | rule impI)+) + fix x1 x1a x2 + assume "x1 = (x1a, x2)" + and "local.stmt (BLOCK ((id0, tp), Some ex') stm) ev cd st = Normal x1" + then have "local.stmt (BLOCK ((id0, tp), Some ex') stm) ev cd st = Normal (x1a, x2)" by simp + then show "gas x2 \ gas st \ P x2" + proof (cases rule: blockSome) + case (1 v t g cd' mem' sck' e') + moreover have "g \ gas st - costs (BLOCK ((id0, tp), Some ex') stm) ev cd st" using msel_ssel_expr_load_rexp_gas(3)[OF 1(2)] by simp + ultimately show ?thesis using assms(1)[OF 1(2)] unfolding wpS_def wp_def by auto + qed +next + fix e + assume "local.stmt (BLOCK ((id0, tp), Some ex') stm) ev cd st = Exception e" + show "E e" using assms(2,3) by (induction rule: Ex.induct) +qed + +end + +subsection \External method invocation\ + +locale Calculus = statement_with_gas + + fixes cname::Identifier + and members:: "(Identifier, Member) fmap" + and const::"(Identifier \ Type) list \ S" + and fb :: S +assumes C1: "ep $$ cname = Some (members, const, fb)" +begin + +text \ + The rules for method invocations is provided in the context of four parameters: + \<^item> @{term_type cname}: The name of the contract to be verified + \<^item> @{term_type members}: The member variables of the contract to be verified + \<^item> @{term const}: The constructor of the contract to be verified + \<^item> @{term fb}: The fallback method of the contract to be verified + +In addition @{thm[source] C1} assigns members, constructor, and fallback method to the contract address. +\ + +text \ + An invariant is a predicate over two parameters: + \<^item> The private store of the contract + \<^item> The balance of the contract +\ + +type_synonym Invariant = "StorageT \ int \ bool" + +subsection \Method invocations and transfer\ + +definition Qe + where "Qe ad iv st \ + (\mid fp f ev. + members $$ mid = Some (Method (fp,True,f)) \ + address ev \ ad + \ (\adex cd st' xe val g v t g' v' e\<^sub>l cd\<^sub>l k\<^sub>l m\<^sub>l g'' acc. + g'' \ gas st \ + type (acc ad) = Some (Contract cname) \ + expr adex ev cd (st'\gas := gas st' - costs (EXTERNAL adex mid xe val) ev cd st'\) (gas st' - costs (EXTERNAL adex mid xe val) ev cd st') = Normal ((KValue ad, Value TAddr), g) \ + expr val ev cd (st'\gas := g\) g = Normal ((KValue v, Value t), g') \ + convert t (TUInt 256) v = Some v' \ + load True fp xe (ffold (init members) (emptyEnv ad cname (address ev) v') (fmdom members)) emptyStore emptyStore emptyStore ev cd (st'\gas := g'\) g' = Normal ((e\<^sub>l, cd\<^sub>l, k\<^sub>l, m\<^sub>l), g'') \ + transfer (address ev) ad v' (accounts (st'\gas := g''\)) = Some acc \ + iv (storage st' ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc ad)) - ReadL\<^sub>i\<^sub>n\<^sub>t v') + \ wpS (stmt f e\<^sub>l cd\<^sub>l) (\st. iv (storage st ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st ad)))) (\e. e = Gas \ e = Err) (st'\gas := g'', accounts:= acc, stack := k\<^sub>l, memory := m\<^sub>l\)))" + +definition Qi + where "Qi ad pre post st \ + (\mid fp f ev. + members $$ mid = Some (Method (fp,False,f)) \ + address ev = ad + \ (\cd st' i xe e\<^sub>l cd\<^sub>l k\<^sub>l m\<^sub>l g. + g \ gas st \ + load False fp xe (ffold (init members) (emptyEnv ad cname (sender ev) (svalue ev)) (fmdom members)) emptyStore emptyStore (memory st') ev cd (st'\gas := gas st' - costs (INVOKE i xe) ev cd st'\) (gas st' - costs (INVOKE i xe) ev cd st') = Normal ((e\<^sub>l, cd\<^sub>l, k\<^sub>l, m\<^sub>l), g) \ + pre mid (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st' ad)), storage st' ad, e\<^sub>l, cd\<^sub>l, k\<^sub>l, m\<^sub>l) + \ wpS (stmt f e\<^sub>l cd\<^sub>l) (\st. post mid (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st ad)), storage st ad)) (\e. e = Gas \ e = Err) (st'\gas := g, stack := k\<^sub>l, memory := m\<^sub>l\)))" + +definition Qfi + where "Qfi ad pref postf st \ + (\ev. address ev = ad + \ (\ex cd st' adex cc v t g g' v' acc. + g' \ gas st \ + expr adex ev cd (st'\gas := gas st' - cc\) (gas st' - cc) = Normal ((KValue ad, Value TAddr), g) \ + expr ex ev cd (st'\gas := g\) g= Normal ((KValue v, Value t), g') \ + convert t (TUInt 256) v = Some v' \ + transfer (address ev) ad v' (accounts st') = Some acc \ + pref (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc ad)), storage st' ad) + \ wpS (\s. stmt fb (ffold (init members) (emptyEnv ad cname (address ev) v') (fmdom members)) emptyStore s) (\st. postf (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st ad)), storage st ad)) (\e. e = Gas \ e = Err) (st'\gas := g', accounts := acc, stack:=emptyStore, memory:=emptyStore\)))" + +definition Qfe + where "Qfe ad iv st \ + (\ev. address ev \ ad + \ (\ex cd st' adex cc v t g g' v' acc. + g' \ gas st \ + type (acc ad) = Some (Contract cname) \ + expr adex ev cd (st'\gas := gas st' - cc\) (gas st' - cc) = Normal ((KValue ad, Value TAddr), g) \ + expr ex ev cd (st'\gas := g\) g = Normal ((KValue v, Value t), g') \ + convert t (TUInt 256) v = Some v' \ + transfer (address ev) ad v' (accounts st') = Some acc \ + iv (storage st' ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc ad)) - ReadL\<^sub>i\<^sub>n\<^sub>t v') + \ wpS (\s. stmt fb (ffold (init members) (emptyEnv ad cname (address ev) v') (fmdom members)) emptyStore s) (\st. iv (storage st ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st ad)))) (\e. e = Gas \ e = Err) (st'\gas := g', accounts := acc, stack:=emptyStore, memory:=emptyStore\)))" + +lemma safeStore[rule_format]: + fixes ad iv + defines "aux1 st \ \st'::State. gas st' < gas st \ Qe ad iv st'" + and "aux2 st \ \st'::State. gas st' < gas st \ Qfe ad iv st'" + shows "\st'. address ev \ ad \ type (accounts st ad) = Some (Contract cname) \ iv (storage st ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st ad))) \ + stmt f ev cd st = Normal ((), st') \ aux1 st \ aux2 st + \ iv (storage st' ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st' ad)))" +proof (induction rule:stmt.induct[where ?P="\f ev cd st. \st'. address ev \ ad \ type (accounts st ad) = Some (Contract cname) \ iv (storage st ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st ad))) \ stmt f ev cd st = Normal ((), st') \ aux1 st \ aux2 st \ iv (storage st' ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st' ad)))"]) + case (1 ev cd st) + show ?case + proof (rule allI, rule impI, (erule conjE)+) + fix st' assume "iv (storage st ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st ad)))" and *: "local.stmt SKIP ev cd st = Normal ((), st')" + then show "iv (storage st' ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st' ad)))" using skip[OF *] by simp + qed +next + case (2 lv ex ev cd st) + show ?case + proof (rule allI, rule impI, (erule conjE)+) + fix st' assume "address ev \ ad" and "iv (storage st ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st ad)))" and 3: "stmt (ASSIGN lv ex) ev cd st = Normal ((), st')" + then show "iv (storage st' ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st' ad)))" by (cases rule: assign[OF 3]; simp) + qed +next + case (3 s1 s2 ev cd st) + show ?case + proof (rule allI, rule impI, (erule conjE)+) + fix st' assume l1: "address ev \ ad" and l12:"type (accounts st ad) = Some (Contract cname)" and l2: "iv (storage st ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st ad)))" and l3: "stmt (COMP s1 s2) ev cd st = Normal ((), st')" and 4:"aux1 st" and 5:"aux2 st" + then show "iv (storage st' ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st' ad)))" + proof (cases rule: comp[OF l3]) + case (1 st'') + moreover have *:"assert Gas (\st. costs (COMP s1 s2) ev cd st < gas st) st = Normal ((), st)" using l3 by (simp add:stmt.simps split:if_split_asm) + moreover from l2 have "iv (storage (st\gas := gas st - costs (COMP s1 s2) ev cd st\) ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts (st\gas := gas st - costs (COMP s1 s2) ev cd st\) ad)))" by simp + moreover have **:"gas (st\gas:= gas st - costs (COMP s1 s2) ev cd st\) \ gas st" by simp + then have "aux1 (st\gas:= gas st - costs (COMP s1 s2) ev cd st\)" using 4 unfolding aux1_def using all_gas_less[OF _ **,of "\st. Qe ad iv st"] by blast + moreover have "aux2 (st\gas:= gas st - costs (COMP s1 s2) ev cd st\)" using 5 unfolding aux2_def using all_gas_less[OF _ **,of "\st. Qfe ad iv st"] by blast + ultimately have "iv (storage st'' ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st'' ad)))" using 3(1) C1 l1 l12 by auto + moreover from l12 have "type (accounts st'' ad) = Some (Contract cname)" using atype_same[OF 1(2), of ad "Contract cname"] l12 by auto + moreover have **:"gas st'' \ gas st" using stmt_gas[OF 1(2)] by simp + then have "aux1 st''" using 4 unfolding aux1_def using all_gas_less[OF _ **,of "\st. Qe ad iv st"] by blast + moreover have "aux2 st''" using 5 unfolding aux2_def using all_gas_less[OF _ **,of "\st. Qfe ad iv st"] by blast + ultimately show ?thesis using 3(2)[OF * _ 1(2), of "()"] 1(3) C1 l1 by simp + qed + qed +next + case (4 ex s1 s2 ev cd st) + show ?case + proof (rule allI, rule impI, (erule conjE)+) + fix st' assume l1: "address ev \ ad" and l12:"type (accounts st ad) = Some (Contract cname)" and l2: "iv (storage st ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st ad)))" and l3: "stmt (ITE ex s1 s2) ev cd st = Normal ((), st')" and l4:"aux1 st" and l5:"aux2 st" + then show "iv (storage st' ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st' ad)))" + proof (cases rule: ite[OF l3]) + case (1 g) + moreover from l2 have "iv (storage (st\gas :=g\) ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts (st\gas := g\) ad)))" by simp + moreover from l12 have "type (accounts (st\gas:= g\) ad) = Some (Contract cname)" using atype_same[OF 1(3), of ad "Contract cname"] l12 by auto + moreover have **:"gas (st\gas:= g\) \ gas st" using msel_ssel_expr_load_rexp_gas(3)[OF 1(2)] by simp + then have "aux1 (st\gas:= g\)" using l4 unfolding aux1_def using all_gas_less[OF _ **,of "\st. Qe ad iv st"] by blast + moreover have "aux2 (st\gas:= g\)" using l5 unfolding aux2_def using all_gas_less[OF _ **,of "\st. Qfe ad iv st"] by blast + ultimately show ?thesis using 4(1) l1 by simp + next + case (2 g) + moreover from l2 have "iv (storage (st\gas := g\) ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts (st\gas := g\) ad)))" by simp + moreover from l12 have "type (accounts (st\gas:= g\) ad) = Some (Contract cname)" using atype_same[OF 2(3), of ad "Contract cname"] l12 by auto + moreover have **:"gas (st\gas:= g\) \ gas st" using msel_ssel_expr_load_rexp_gas(3)[OF 2(2)] by simp + then have "aux1 (st\gas:= g\)" using l4 unfolding aux1_def using all_gas_less[OF _ **,of "\st. Qe ad iv st"] by blast + moreover have "aux2 (st\gas:= g\)" using l5 unfolding aux2_def using all_gas_less[OF _ **,of "\st. Qfe ad iv st"] by blast + ultimately show ?thesis using 4(2) l1 true_neq_false by simp + qed + qed +next + case (5 ex s0 ev cd st) + show ?case + proof (rule allI, rule impI, (erule conjE)+) + fix st' assume l1: "address ev \ ad" and l12:"type (accounts st ad) = Some (Contract cname)" and l2: "iv (storage st ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st ad)))" and l3: "stmt (WHILE ex s0) ev cd st = Normal ((), st')" and l4:"aux1 st" and l5:"aux2 st" + then show "iv (storage st' ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st' ad)))" + proof (cases rule: while[OF l3]) + case (1 g st'') + moreover from l2 have "iv (storage (st\gas :=g\) ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts (st\gas := g\) ad)))" by simp + moreover have *:"gas (st\gas:= g\) \ gas st" using msel_ssel_expr_load_rexp_gas(3)[OF 1(2)] by simp + then have "aux1 (st\gas:= g\)" using l4 unfolding aux1_def using all_gas_less[OF _ *,of "\st. Qe ad iv st"] by blast + moreover have "aux2 (st\gas:= g\)" using l5 unfolding aux2_def using all_gas_less[OF _ *,of "\st. Qfe ad iv st"] by blast + ultimately have "iv (storage st'' ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st'' ad)))" using 5(1) l1 l12 by simp + moreover from l12 have "type (accounts st'' ad) = Some (Contract cname)" using atype_same[OF 1(3), of ad "Contract cname"] l12 by auto + moreover have **:"gas st'' \ gas st" using stmt_gas[OF 1(3)] * by simp + then have "aux1 st''" using l4 unfolding aux1_def using all_gas_less[OF _ **,of "\st. Qe ad iv st"] by blast + moreover have "aux2 st''" using l5 unfolding aux2_def using all_gas_less[OF _ **,of "\st. Qfe ad iv st"] by blast + ultimately show ?thesis using 5(2) 1(1,2,3,4) l1 by simp + next + case (2 g) + then show "iv (storage st' ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st' ad)))" using l2 by simp + qed + qed +next + case (6 i xe ev cd st) + show ?case + proof (rule allI, rule impI, (erule conjE)+) + fix st' assume 1: "address ev \ ad" and l12:"type (accounts st ad) = Some (Contract cname)" and 2: "iv (storage st ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st ad)))" and 3: "local.stmt (INVOKE i xe) ev cd st = Normal ((), st')" and 4:"aux1 st" and 5:"aux2 st" + from 3 obtain ct fb' fp f e\<^sub>l cd\<^sub>l k\<^sub>l m\<^sub>l g st'' + where l0: "costs (INVOKE i xe) ev cd st < gas st" + and l1: "ep $$ contract ev = Some (ct, fb')" + and l2: "ct $$ i = Some (Method (fp, False, f))" + and l3: "load False fp xe (ffold (init ct) (emptyEnv (address ev) (contract ev) (sender ev) (svalue ev)) (fmdom ct)) emptyStore emptyStore (memory (st\gas := gas st - costs (INVOKE i xe) ev cd st\)) ev cd (st\gas := gas st - costs (INVOKE i xe) ev cd st\) (gas st - costs (INVOKE i xe) ev cd st) = Normal ((e\<^sub>l, cd\<^sub>l, k\<^sub>l, m\<^sub>l), g)" + and l4: "stmt f e\<^sub>l cd\<^sub>l (st\gas:= g, stack:=k\<^sub>l, memory:=m\<^sub>l\) = Normal ((), st'')" + and l5: "st' = st''\stack:=stack st\" + using invoke by blast + from 3 have *:"assert Gas (\st. costs (INVOKE i xe) ev cd st < gas st) st = Normal ((), st)" by (simp add:stmt.simps split:if_split_asm) + moreover have **: "modify (\st. st\gas := gas st - costs (INVOKE i xe) ev cd st\) st = Normal ((), st\gas := gas st - costs (INVOKE i xe) ev cd st\)" by simp + ultimately have "\st'. address e\<^sub>l \ ad \ iv (storage (st\gas := g, stack := k\<^sub>l, memory := m\<^sub>l\) ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts (st\gas := g, stack := k\<^sub>l, memory := m\<^sub>l\) ad))) \ local.stmt f e\<^sub>l cd\<^sub>l (st\gas := g, stack := k\<^sub>l, memory := m\<^sub>l\) = Normal ((), st') \ aux1 (st\gas := g, stack := k\<^sub>l, memory := m\<^sub>l\) \ aux2 (st\gas := g, stack := k\<^sub>l, memory := m\<^sub>l\) \ + iv (storage st' ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st' ad)))" using 6[OF * **] l1 l2 l3 l12 by (simp split:if_split_asm result.split_asm prod.split_asm option.split_asm) + moreover have "address (ffold (init ct) (emptyEnv (address ev) (contract ev) (sender ev) (svalue ev)) (fmdom ct)) = address ev" using ffold_init_ad_same[of ct "(emptyEnv (address ev) (contract ev) (sender ev) (svalue ev))"] unfolding emptyEnv_def by simp + with 1 have "address e\<^sub>l \ ad" using msel_ssel_expr_load_rexp_gas(4)[OF l3] by simp + moreover from 2 have "iv (storage (st\gas:= g, stack:=k\<^sub>l, memory:=m\<^sub>l\) ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts (st\gas:= g, stack:=k\<^sub>l, memory:=m\<^sub>l\) ad)))" by simp + moreover have *:"gas (st\gas := g, stack := k\<^sub>l, memory := m\<^sub>l\) \ gas st" using msel_ssel_expr_load_rexp_gas(4)[OF l3] by auto + then have "aux1 (st\gas:= g, stack:=k\<^sub>l, memory:=m\<^sub>l\)" using 4 unfolding aux1_def using all_gas_less[OF _ *,of "\st. Qe ad iv st"] by blast + moreover have *:"gas (st\gas := g, stack := k\<^sub>l, memory := m\<^sub>l\) \ gas st" using msel_ssel_expr_load_rexp_gas(4)[OF l3] by auto + then have "aux2 (st\gas := g, stack := k\<^sub>l, memory := m\<^sub>l\)" using 5 unfolding aux2_def using all_gas_less[OF _ *,of "\st. Qfe ad iv st"] by blast + ultimately have "iv (storage st'' ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st'' ad)))" using l4 C1 by auto + then show "iv (storage st' ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st' ad)))" using l5 by simp + qed +next + case (7 ad' i xe val ev cd st) + show ?case + proof (rule allI, rule impI, (erule conjE)+) + fix st' assume l1: "address ev \ ad" + and l12:"type (accounts st ad) = Some (Contract cname)" + and l2: "iv (storage st ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st ad)))" + and 3: "local.stmt (EXTERNAL ad' i xe val) ev cd st = Normal ((), st')" + and 4: "aux1 st" and 5:"aux2 st" + show "iv (storage st' ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st' ad)))" + proof (cases rule: external[OF 3]) + case (1 adv c g ct cn fb' v t g' v' fp f e\<^sub>l cd\<^sub>l k\<^sub>l m\<^sub>l g'' acc st'') + then show ?thesis + proof (cases "adv = ad") + case True + moreover from this have "cname = c" using l12 1(4) by simp + moreover from this have "members = ct" using C1 1(5) by simp + moreover have "gas st \ costs (EXTERNAL ad' i xe val) ev cd st" using 3 by (simp add:stmt.simps split:if_split_asm) + then have "g'' < gas st" using msel_ssel_expr_load_rexp_gas(3)[OF 1(2)] msel_ssel_expr_load_rexp_gas(3)[OF 1(6)] msel_ssel_expr_load_rexp_gas(4)[OF 1(9)] external_not_zero[of ad' i xe val ev cd st] by auto + then have "Qe ad iv (st\gas := g'', accounts := acc, stack := k\<^sub>l, memory := m\<^sub>l\)" using 4 unfolding aux1_def by simp + moreover have "g'' \ gas (st\gas := g'', accounts := acc, stack := k\<^sub>l, memory := m\<^sub>l\)" by simp + moreover from l12 have "type (acc ad) = Some (Contract cname)" using transfer_type_same[OF 1(10)] by simp + moreover have "i |\| fmdom members" using 1(8) `members = ct` by (simp add: fmdomI) + moreover have "members $$ i = Some (Method (fp,True,f))" using 1(8) `members = ct` by simp + moreover have "iv (storage st ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc ad)) - ReadL\<^sub>i\<^sub>n\<^sub>t v')" + proof - + have "iv (storage st ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st ad)))" using l2 by simp + moreover have "ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc ad)) = ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st ad)) + ReadL\<^sub>i\<^sub>n\<^sub>t v'" using transfer_add[OF 1(10)] l1 True by simp + ultimately show ?thesis by simp + qed + ultimately have "wpS (local.stmt f e\<^sub>l cd\<^sub>l) (\st. iv (storage st ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st ad)))) (\e. e = Gas \ e = Err) (st\gas := g'', accounts := acc, stack := k\<^sub>l, memory := m\<^sub>l\)" unfolding Qe_def using l1 l12 1(2) 1(6-10) by blast + moreover have "stmt f e\<^sub>l cd\<^sub>l (st\gas := g'', accounts := acc, stack := k\<^sub>l, memory := m\<^sub>l\) = Normal ((), st'')" using 1(11) by simp + ultimately show "iv (storage st' ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st' ad)))" unfolding wpS_def wp_def using 1(12) by simp + next + case False + + from 3 have *:"assert Gas (\st. costs (EXTERNAL ad' i xe val) ev cd st < gas st) st = Normal ((), st)" by (simp add:stmt.simps split:if_split_asm) + moreover have **: "modify (\st. st\gas := gas st - costs (EXTERNAL ad' i xe val) ev cd st\) st = Normal ((), st\gas := gas st - costs (EXTERNAL ad' i xe val) ev cd st\)" by simp + ultimately have "\st'. address e\<^sub>l \ ad \ type (acc ad) = Some (Contract cname) \ iv (storage st ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc ad))) \ local.stmt f e\<^sub>l cd\<^sub>l (st\gas := g'', accounts := acc, stack := k\<^sub>l, memory := m\<^sub>l\) = Normal ((), st') \ aux1 (st\gas := g'', accounts := acc, stack := k\<^sub>l, memory := m\<^sub>l\) \ aux2 (st\gas := g'', accounts := acc, stack := k\<^sub>l, memory := m\<^sub>l\) \ iv (storage st' ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st' ad)))" using 7(1)[OF * **] 1 by simp + moreover have "address (ffold (init ct) (emptyEnv adv c (address ev) v') (fmdom ct)) = adv" using ffold_init_ad_same[of ct "(emptyEnv adv c (address ev) v)"] unfolding emptyEnv_def by simp + with False have "address e\<^sub>l \ ad" using msel_ssel_expr_load_rexp_gas(4)[OF 1(9)] by simp + moreover have "bal (acc ad) = bal ((accounts st) ad)" using transfer_eq[OF 1(10)] False l1 by simp + then have "iv (storage st ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc ad)))" using l2 by simp + moreover have "type (acc ad) = Some (Contract cname)" using transfer_type_same l12 1(10) by simp + moreover have *:"gas (st\gas := g'', accounts := acc, stack := k\<^sub>l, memory := m\<^sub>l\) \ gas st" using msel_ssel_expr_load_rexp_gas(3)[OF 1(2)] msel_ssel_expr_load_rexp_gas(3)[OF 1(6)] msel_ssel_expr_load_rexp_gas(4)[OF 1(9)] by auto + then have "aux1 (st\gas := g'', accounts := acc, stack := k\<^sub>l, memory := m\<^sub>l\)" using 4 unfolding aux1_def using all_gas_less[OF _ *,of "\st. Qe ad iv st"] by blast + moreover have "aux2 (st\gas := g'', accounts := acc, stack := k\<^sub>l, memory := m\<^sub>l\)" using 5 unfolding aux2_def using all_gas_less[OF _ *,of "\st. Qfe ad iv st"] by blast + ultimately have "iv (storage st'' ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st'' ad)))" using 1(11) by simp + then show "iv (storage st' ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st' ad)))" using 1(12) by simp + qed + next + case (2 adv c g ct cn fb' v t g' v' acc st'') + then show ?thesis + proof (cases "adv = ad") + case True + moreover have "gas st \ costs (EXTERNAL ad' i xe val) ev cd st" using 3 by (simp add:stmt.simps split:if_split_asm) + then have "gas (st\gas := g', accounts := acc, stack := emptyStore, memory := emptyStore\) < gas st" using msel_ssel_expr_load_rexp_gas(3)[OF 2(2)] msel_ssel_expr_load_rexp_gas(3)[OF 2(6)] external_not_zero[of ad' i xe val ev cd st] by simp + then have "Qfe ad iv (st\gas := g', accounts := acc, stack := emptyStore, memory := emptyStore\)" using 5 unfolding aux2_def by simp + moreover have "iv (storage st ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc ad)) - ReadL\<^sub>i\<^sub>n\<^sub>t v')" + proof - + have "iv (storage st ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st ad)))" using l2 by simp + moreover have "ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc ad)) = ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st ad)) + ReadL\<^sub>i\<^sub>n\<^sub>t v'" using transfer_add[OF 2(9)] l1 True by simp + ultimately show ?thesis by simp + qed + moreover have "g' \ gas (st\gas := g', accounts := acc, stack := emptyStore, memory := emptyStore\)" by simp + moreover from l12 have "type (acc ad) = Some (Contract cname)" using transfer_type_same[OF 2(9)] by simp + ultimately have "wpS (local.stmt fb (ffold (init members) (emptyEnv ad cname (address ev) v') (fmdom members)) emptyStore) (\st. iv (storage st ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st ad)))) (\e. e = Gas \ e = Err) (st\gas := g', accounts := acc, stack := emptyStore, memory := emptyStore\)" unfolding Qfe_def using l1 l12 2(2) 2(6-9) by blast + moreover have "stmt fb (ffold (init members) (emptyEnv ad cname (address ev) v') (fmdom members)) emptyStore (st\gas := g', accounts := acc, stack := emptyStore, memory := emptyStore\) = Normal ((), st'')" + proof - + from True have "cname = c" using l12 2(4) by simp + moreover from this have "fb'=fb" using C1 2(5) by simp + moreover from True `cname = c` have "members = ct" using C1 2(5) by simp + ultimately show ?thesis using 2(10) True by simp + qed + ultimately show "iv (storage st' ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st' ad)))" unfolding wpS_def wp_def using 2(11) by simp + next + case False + + from 3 have *:"assert Gas (\st. costs (EXTERNAL ad' i xe val) ev cd st < gas st) st = Normal ((), st)" by (simp add:stmt.simps split:if_split_asm) + then have "\st'. address (ffold (init ct) (emptyEnv adv c (address ev) v) (fmdom ct)) \ ad \ + type (acc ad) = Some (Contract cname) \ + iv (storage st ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc ad))) \ + local.stmt fb' (ffold (init ct) (emptyEnv adv c (address ev) v') (fmdom ct)) emptyStore (st\gas := g', accounts := acc, stack := emptyStore, memory := emptyStore\) = Normal ((), st') \ aux1 (st\gas := g', accounts := acc, stack := emptyStore, memory := emptyStore\) \ aux2 (st\gas := g', accounts := acc, stack := emptyStore, memory := emptyStore\) + \ iv (storage st' ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st' ad)))" using 7(2)[OF *] 2 by simp + moreover from False have "address (ffold (init ct) (emptyEnv adv c (address ev) v') (fmdom ct)) \ ad" using ffold_init_ad_same[where ?e="\address = adv, contract = c, sender = address ev, svalue = v, denvalue = {$$}\" and ?e'="ffold (init ct) (emptyEnv adv c (address ev) v) (fmdom ct)"] unfolding emptyEnv_def by simp + moreover have "bal (acc ad) = bal ((accounts st) ad)" using transfer_eq[OF 2(9)] False l1 by simp + then have "iv (storage st ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc ad)))" + using l2 by simp + moreover have "type (acc ad) = Some (Contract cname)" using transfer_type_same l12 2(9) by simp + moreover have *:"gas (st\gas := g', accounts := acc, stack := emptyStore, memory := emptyStore\) \ gas st" using msel_ssel_expr_load_rexp_gas(3)[OF 2(2)] msel_ssel_expr_load_rexp_gas(3)[OF 2(6)] by simp + then have "aux1 (st\gas := g', accounts := acc, stack := emptyStore, memory := emptyStore\)" using 4 unfolding aux1_def using all_gas_less[OF _ *,of "\st. Qe ad iv st"] by blast + moreover have "aux2 (st\gas := g', accounts := acc, stack := emptyStore, memory := emptyStore\)" using 5 unfolding aux2_def using all_gas_less[OF _ *,of "\st. Qfe ad iv st"] by blast + ultimately have "iv (storage st'' ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st'' ad)))" using 2(10) by simp + then show "iv (storage st' ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st' ad)))" using 2(11) by simp + qed + qed + qed +next + case (8 ad' ex ev cd st) + show ?case + proof (rule allI, rule impI, (erule conjE)+) + fix st' assume l1: "address ev \ ad" and l12:"type (accounts st ad) = Some (Contract cname)" and l2: "iv (storage st ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st ad)))" and 3: "local.stmt (TRANSFER ad' ex) ev cd st = Normal ((), st')" and 4:"aux1 st" and 5:"aux2 st" + show "iv (storage st' ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st' ad)))" + proof (cases rule: transfer[OF 3]) + case (1 v t g adv c g' v' acc ct cn f st'') + then show ?thesis + proof (cases "adv = ad") + case True + moreover have "gas st \ costs (TRANSFER ad' ex) ev cd st" using 3 by (simp add:stmt.simps split:if_split_asm) + then have "gas (st\gas := g', accounts := acc, stack := emptyStore, memory := emptyStore\) < gas st" using msel_ssel_expr_load_rexp_gas(3)[OF 1(2)] msel_ssel_expr_load_rexp_gas(3)[OF 1(3)] transfer_not_zero[of ad' ex ev cd st] by auto + then have "Qfe ad iv (st\gas := g', accounts := acc, stack := emptyStore, memory := emptyStore\)" using 5 unfolding aux2_def by simp + moreover have "sender (ffold (init ct) (emptyEnv adv c (address ev) v') (fmdom ct)) \ ad" using l1 ffold_init_ad_same[where ?e = "(emptyEnv adv c (address ev) v')" and ?e'="ffold (init ct) (emptyEnv adv c (address ev) v') (fmdom ct)"] unfolding emptyEnv_def by simp + moreover have "svalue (ffold (init ct) (emptyEnv adv c (address ev) v') (fmdom ct)) = v'" using ffold_init_ad_same[where ?e = "(emptyEnv adv c (address ev) v')" and ?e'="ffold (init ct) (emptyEnv adv c (address ev) v') (fmdom ct)", of ct "fmdom ct"] unfolding emptyEnv_def by simp + moreover have "iv (storage st ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc ad)) - ReadL\<^sub>i\<^sub>n\<^sub>t v')" + proof - + have "iv (storage st ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st ad)))" using l2 by simp + moreover have "ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc ad)) = ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st ad)) + ReadL\<^sub>i\<^sub>n\<^sub>t v'" using transfer_add[OF 1(7)] l1 True by simp + ultimately show ?thesis by simp + qed + moreover have "g' \ gas (st\gas := g', accounts := acc, stack := emptyStore, memory := emptyStore\)" by simp + moreover from l12 have "type (acc ad) = Some (Contract cname)" using transfer_type_same[OF 1(7)] by simp + ultimately have "wpS (local.stmt fb (ffold (init members) (emptyEnv ad cname (address ev) v') (fmdom members)) emptyStore) (\st. iv (storage st ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st ad)))) (\e. e = Gas \ e = Err) (st\gas := g', accounts := acc, stack := emptyStore, memory := emptyStore\)" unfolding Qfe_def using l1 l12 1(2-7) by blast + moreover have "stmt fb (ffold (init members) (emptyEnv ad cname (address ev) v') (fmdom members)) emptyStore (st\gas := g', accounts := acc, stack := emptyStore, memory := emptyStore\) = Normal ((), st'')" + proof - + from True have "cname = c" using l12 1(5) by simp + moreover from this have "f=fb" using C1 1(6) by simp + moreover from True `cname = c` have "members = ct" using C1 1(6) by simp + ultimately show ?thesis using 1(8) True by simp + qed + ultimately show "iv (storage st' ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st' ad)))" unfolding wpS_def wp_def using 1(9) by simp + next + case False + + from 3 have *:"assert Gas (\st. costs (TRANSFER ad' ex) ev cd st < gas st) st = Normal ((), st)" by (simp add:stmt.simps split:if_split_asm) + then have "\st'. address (ffold (init ct) (emptyEnv adv c (address ev) v) (fmdom ct)) \ ad \ + type (acc ad) = Some (Contract cname) \ + iv (storage st ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc ad))) \ + local.stmt f (ffold (init ct) (emptyEnv adv c (address ev) v') (fmdom ct)) emptyStore (st\gas := g', accounts := acc, stack := emptyStore, memory := emptyStore\) = Normal ((), st') \ + aux1 (st\gas := g', accounts := acc, stack := emptyStore, memory := emptyStore\) \ aux2 (st\gas := g', accounts := acc, stack := emptyStore, memory := emptyStore\) + \ iv (storage st' ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st' ad)))" using 8(1)[OF *] 1 by simp + moreover from False have "address (ffold (init ct) (emptyEnv adv c (address ev) v') (fmdom ct)) \ ad" using ffold_init_ad_same[of ct "emptyEnv adv c (address ev) v"] unfolding emptyEnv_def by simp + moreover have "bal (acc ad) = bal ((accounts st) ad)" using transfer_eq[OF 1(7)] False l1 by simp + then have "iv (storage st ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc ad)))" + using l2 by simp + moreover have "type (acc ad) = Some (Contract cname)" using transfer_type_same l12 1(7) by simp + moreover have *:"gas (st\gas := g', accounts := acc, stack := emptyStore, memory := emptyStore\) \ gas st" using msel_ssel_expr_load_rexp_gas(3)[OF 1(2)] msel_ssel_expr_load_rexp_gas(3)[OF 1(3)] by simp + then have "aux1 (st\gas := g', accounts := acc, stack := emptyStore, memory := emptyStore\)" using 4 unfolding aux1_def using all_gas_less[OF _ *,of "\st. Qe ad iv st"] by blast + moreover have "aux2 (st\gas := g', accounts := acc, stack := emptyStore, memory := emptyStore\)" using 5 unfolding aux2_def using all_gas_less[OF _ *,of "\st. Qfe ad iv st"] by blast + ultimately have "iv (storage st'' ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st'' ad)))" using 1(8) C1 by simp + then show "iv (storage st' ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st' ad)))" using 1(9) by simp + qed + next + case (2 v t g adv g' v' acc) + moreover from 2(5) have "adv \ ad" using C1 l12 by auto + then have "bal (acc ad) = bal (accounts st ad)" using transfer_eq[OF 2(6)] l1 by simp + ultimately show ?thesis using l2 by simp + qed + qed +next + case (9 id0 tp s ev cd st) + show ?case + proof (rule allI, rule impI, (erule conjE)+) + fix st' assume l1: "address ev \ ad" and l12:"type (accounts st ad) = Some (Contract cname)" and l2: "iv (storage st ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st ad)))" and l3: "stmt (BLOCK ((id0, tp), None) s) ev cd st = Normal ((), st')" and 4:"aux1 st" and 5:"aux2 st" + then show "iv (storage st' ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st' ad)))" + proof (cases rule: blockNone[OF l3]) + case (1 cd' mem' sck' e') + moreover from l2 have "iv (storage (st\gas := gas st - costs (BLOCK ((id0, tp), None) s) ev cd st, stack := sck', memory := mem'\) ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts (st\gas := gas st - costs (BLOCK ((id0, tp), None) s) ev cd st, stack := sck', memory := mem'\) ad)))" by simp + moreover have *:"gas (st\gas:= gas st - costs (BLOCK ((id0, tp), None) s) ev cd st, stack := sck', memory := mem'\) \ gas st" by simp + then have "aux1 (st\gas:= gas st - costs (BLOCK ((id0, tp), None) s) ev cd st, stack := sck', memory := mem'\)" using 4 unfolding aux1_def using all_gas_less[OF _ *,of "\st. Qe ad iv st"] by blast + moreover have "aux2 (st\gas:= gas st - costs (BLOCK ((id0, tp), None) s) ev cd st, stack := sck', memory := mem'\)" using 5 unfolding aux2_def using all_gas_less[OF _ *,of "\st. Qfe ad iv st"] by blast + moreover have "address e' \ ad" using decl_env[OF 1(2)] l1 by simp + ultimately show "iv (storage st' ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st' ad)))" using 9(1) l12 by simp + qed + qed +next + case (10 id0 tp ex' s ev cd st) + show ?case + proof (rule allI, rule impI, (erule conjE)+) + fix st' assume l1: "address ev \ ad" and l12:"type (accounts st ad) = Some (Contract cname)" and l2: "iv (storage st ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st ad)))" and l3: "stmt (BLOCK ((id0, tp), Some ex') s) ev cd st = Normal ((), st')" and 4:"aux1 st" and 5:"aux2 st" + then show "iv (storage st' ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st' ad)))" + proof (cases rule: blockSome[OF l3]) + case (1 v t g cd' mem' sck' e') + moreover from l2 have "iv (storage (st\gas := g, stack := sck', memory := mem'\) ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts (st\gas := g, stack := sck', memory := mem'\) ad)))" by simp + moreover have *:"gas (st\gas:= g, stack := sck', memory := mem'\) \ gas st" using msel_ssel_expr_load_rexp_gas(3)[OF 1(2)] by simp + then have "aux1 (st\gas:= g, stack := sck', memory := mem'\)" using 4 unfolding aux1_def using all_gas_less[OF _ *,of "\st. Qe ad iv st"] by blast + moreover have "aux2 (st\gas:= g, stack := sck', memory := mem'\)" using 5 unfolding aux2_def using all_gas_less[OF _ *,of "\st. Qfe ad iv st"] by blast + moreover have "address e' \ ad" using decl_env[OF 1(3)] l1 by simp + ultimately show "iv (storage st' ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st' ad)))" using 10(1) l12 by simp + qed + qed +next + case (11 i xe val ev cd st) + show ?case + proof (rule allI, rule impI, (erule conjE)+) + fix st' assume l1: "address ev \ ad" and l12:"type (accounts st ad) = Some (Contract cname)" and l2: "iv (storage st ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st ad)))" and l3: "stmt (NEW i xe val) ev cd st = Normal ((), st')" and 4:"aux1 st" and 5:"aux2 st" + then show "iv (storage st' ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st' ad)))" + proof (cases rule: new[OF l3]) + case (1 v t g ct cn fb e\<^sub>l cd\<^sub>l k\<^sub>l m\<^sub>l g' acc st'') + moreover define adv where "adv = hash (address ev) \contracts (accounts (st\gas := gas st - costs (NEW i xe val) ev cd st\) (address ev))\" + moreover define st''' where "st''' = (st\gas := gas st - costs (NEW i xe val) ev cd st, gas := g', accounts := (accounts st)(adv := \bal = ShowL\<^sub>i\<^sub>n\<^sub>t 0, type = Some (Contract i), contracts = 0\), storage := (storage st)(adv := {$$}), accounts := acc, stack := k\<^sub>l, memory := m\<^sub>l\)" + ultimately have "\st'. address e\<^sub>l \ ad \ + type (accounts st''' ad) = Some (Contract cname) \ + iv (storage st''' ad) \bal (accounts st''' ad)\ \ + local.stmt (snd cn) e\<^sub>l cd\<^sub>l st''' = Normal ((), st') \ aux1 st''' \ aux2 st''' \ + iv (storage st' ad) \bal (accounts st' ad)\" + using 11 by simp + moreover have "address e\<^sub>l \ ad" + proof - + have "address e\<^sub>l = adv" using msel_ssel_expr_load_rexp_gas(4)[OF 1(5)] adv_def by simp + moreover have "adv \ ad" using l12 1(2) adv_def by auto + ultimately show ?thesis by simp + qed + moreover have "type (accounts st''' ad) = Some (Contract cname)" + proof - + have "adv \ ad" using l12 1(2) adv_def by auto + then have "type (accounts st ad) = type (acc ad)" using transfer_type_same[OF 1(6)] adv_def by simp + then show ?thesis using l12 st'''_def by simp + qed + moreover have "iv (storage st''' ad) \bal (accounts st''' ad)\" + proof - + have "adv \ ad" using l12 1(2) adv_def by auto + then have "bal (accounts st ad) = bal (accounts st''' ad)" using transfer_eq[OF 1(6), of ad] l1 using st'''_def adv_def by simp + moreover have "storage st ad = storage st''' ad" using st'''_def `adv \ ad` by simp + ultimately show ?thesis using l2 by simp + qed + moreover have "local.stmt (snd cn) e\<^sub>l cd\<^sub>l st''' = Normal ((), st'')" using 1(7) st'''_def adv_def by simp + moreover have "aux1 st'''" + proof - + have *: "gas st''' \ gas st" unfolding st'''_def using msel_ssel_expr_load_rexp_gas(3)[OF 1(3)] msel_ssel_expr_load_rexp_gas(4)[OF 1(5)] by auto + then show ?thesis using 4 unfolding aux1_def using all_gas_less[OF _ *,of "\st. Qe ad iv st"] by simp + qed + moreover have "aux2 st'''" + proof - + have *: "gas st''' \ gas st" unfolding st'''_def using msel_ssel_expr_load_rexp_gas(3)[OF 1(3)] msel_ssel_expr_load_rexp_gas(4)[OF 1(5)] by auto + then show ?thesis using 5 unfolding aux2_def using all_gas_less[OF _ *,of "\st. Qe ad iv st"] by simp + qed + ultimately have "iv (storage st'' ad) \bal (accounts st'' ad)\" by simp + moreover have "storage st'' ad = storage st' ad" using 1(8) by simp + moreover have "bal (accounts st'' ad) = bal (accounts st' ad)" using 1(8) by simp + ultimately show ?thesis by simp + qed + qed +qed + +type_synonym Precondition = "int \ StorageT \ Environment \ Memoryvalue Store \ Stackvalue Store \ Memoryvalue Store \ bool" +type_synonym Postcondition = "int \ StorageT \ bool" + +text \ + The following lemma can be used to verify (recursive) internal or external method calls and transfers executed from **inside** (@{term "address ev = ad"}). + In particular the lemma requires the contract to be annotated as follows: + \<^item> Pre/Postconditions for internal methods + \<^item> Invariants for external methods + + The lemma then requires us to verify the following: + \<^item> Postconditions from preconditions for internal method bodies. + \<^item> Invariants hold for external method bodies. + + To this end it allows us to assume the following: + \<^item> Preconditions imply postconditions for internal method calls. + \<^item> Invariants hold for external method calls for other contracts external methods. +\ + +definition Pe + where "Pe ad iv st \ + (\ev ad' i xe val cd. + address ev = ad \ + (\adv c g v t g' v'. + expr ad' ev cd (st\gas := gas st - costs (EXTERNAL ad' i xe val) ev cd st\) (gas st - costs (EXTERNAL ad' i xe val) ev cd st) = Normal ((KValue adv, Value TAddr), g) \ + adv \ ad \ + type (accounts st adv) = Some (Contract c) \ + c |\| fmdom ep \ + expr val ev cd (st\gas := g\) g = Normal ((KValue v, Value t), g') \ + convert t (TUInt 256) v = Some v' + \ iv (storage st ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st ad)) - ReadL\<^sub>i\<^sub>n\<^sub>t v')) + \ wpS (\s. stmt (EXTERNAL ad' i xe val) ev cd s) (\st. iv (storage st ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st ad)))) (\e. e = Gas \ e = Err) st)" + +definition Pi + where "Pi ad pre post st \ + (\ev i xe cd. + address ev = ad \ + contract ev = cname \ + (\fp e\<^sub>l cd\<^sub>l k\<^sub>l m\<^sub>l g. + load False fp xe (ffold (init members) (emptyEnv ad (contract ev) (sender ev) (svalue ev)) (fmdom members)) emptyStore emptyStore (memory st) ev cd (st\gas := gas st - costs (INVOKE i xe) ev cd st\) (gas st - costs (INVOKE i xe) ev cd st) = Normal ((e\<^sub>l, cd\<^sub>l, k\<^sub>l, m\<^sub>l), g) + \ pre i (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st ad)), storage st ad, e\<^sub>l, cd\<^sub>l, k\<^sub>l, m\<^sub>l)) + \ wpS (\s. stmt (INVOKE i xe) ev cd s) (\st. post i (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st ad)), storage st ad)) (\e. e = Gas \ e = Err) st)" + +definition Pfi + where "Pfi ad pref postf st \ + (\ev ex ad' cd. + address ev = ad \ + (\adv g. + expr ad' ev cd (st\gas := gas st - costs (TRANSFER ad' ex) ev cd st\) (gas st - costs (TRANSFER ad' ex) ev cd st) = Normal ((KValue adv, Value TAddr), g) + \ adv = ad) \ + (\g v t g'. + expr ad' ev cd (st\gas := gas st - costs (TRANSFER ad' ex) ev cd st\) (gas st - costs (TRANSFER ad' ex) ev cd st) = Normal ((KValue ad, Value TAddr), g) \ + expr ex ev cd (st\gas := g\) g = Normal ((KValue v, Value t), g') + \ pref (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st ad)), storage st ad)) + \ wpS (\s. stmt (TRANSFER ad' ex) ev cd s) (\st. postf (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st ad)), storage st ad)) (\e. e = Gas \ e = Err) st)" + +definition Pfe + where "Pfe ad iv st \ + (\ev ex ad' cd. + address ev = ad \ + (\adv g. + expr ad' ev cd (st\gas := gas st - costs (TRANSFER ad' ex) ev cd st\) (gas st - costs (TRANSFER ad' ex) ev cd st) = Normal ((KValue adv, Value TAddr), g) + \ adv \ ad) \ + (\adv g v t g' v'. + expr ad' ev cd (st\gas := gas st - costs (TRANSFER ad' ex) ev cd st\) (gas st - costs (TRANSFER ad' ex) ev cd st) = Normal ((KValue adv, Value TAddr), g) \ + adv \ ad \ + expr ex ev cd (st\gas := g\) g = Normal ((KValue v, Value t), g') \ + convert t (TUInt 256) v = Some v' + \ iv (storage st ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st ad)) - ReadL\<^sub>i\<^sub>n\<^sub>t v')) + \ wpS (\s. stmt (TRANSFER ad' ex) ev cd s) (\st. iv (storage st ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st ad)))) (\e. e = Gas \ e = Err) st)" + +lemma wp_external_invoke_transfer: + fixes pre::"Identifier \ Precondition" + and post::"Identifier \ Postcondition" + and pref::"Postcondition" + and postf::"Postcondition" + and iv::"Invariant" + assumes assm: "\st::State. + \\st'::State. gas st' \ gas st \ type (accounts st' ad) = Some (Contract cname) + \ Pe ad iv st' \ Pi ad pre post st' \ Pfi ad pref postf st' \ Pfe ad iv st'\ + \ Qe ad iv st \ Qi ad pre post st \ Qfi ad pref postf st \ Qfe ad iv st" + shows "type (accounts st ad) = Some (Contract cname) \ Pe ad iv st \ Pi ad pre post st \ Pfi ad pref postf st \ Pfe ad iv st" +proof (induction st rule: gas_induct) + case (1 st) + show ?case unfolding Pe_def Pi_def Pfi_def Pfe_def + proof elims + fix ev::Environment and ad' i xe val cd + assume a00: "type (accounts st ad) = Some (Contract cname)" + and a0: "address ev = ad" + and a1: "\adv c g v t g' v'. + local.expr ad' ev cd (st\gas := gas st - costs (EXTERNAL ad' i xe val) ev cd st\) + (gas st - costs (EXTERNAL ad' i xe val) ev cd st) = + Normal ((KValue adv, Value TAddr), g) \ + adv \ ad \ + type (accounts st adv) = Some (Contract c) \ + c |\| fmdom ep \ + local.expr val ev cd (st\gas := g\) g = Normal ((KValue v, Value t), g') \ + convert t (TUInt 256) v = Some v' + \ iv (storage st ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st ad)) - ReadL\<^sub>i\<^sub>n\<^sub>t v')" + show "wpS (local.stmt (EXTERNAL ad' i xe val) ev cd) (\st. iv (storage st ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st ad)))) (\e. e = Gas \ e = Err) st" unfolding wpS_def wp_def + proof (split result.split; split prod.split; rule conjI; (rule allI)+; (rule impI)+) + fix x1 x1a s'''''' + assume "x1 = (x1a, s'''''')" and 2: "local.stmt (EXTERNAL ad' i xe val) ev cd st = Normal x1" + then have "local.stmt (EXTERNAL ad' i xe val) ev cd st = Normal (x1a, s'''''')" by simp + then show "gas s'''''' \ gas st \ iv (storage s'''''' ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts s'''''' ad)))" + proof (cases rule: external) + case (Some adv0 c0 g0 ct0 cn0 fb0 v0 t0 g0' v0' fp0 f0 e\<^sub>l0 cd\<^sub>l0 k\<^sub>l0 m\<^sub>l0 g0'' acc0 st0'') + moreover have "iv (storage st0'' ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st0'' ad)))" + proof - + from Some(3) have "adv0 \ ad" using a0 by simp + then have "address e\<^sub>l0 \ ad" using msel_ssel_expr_load_rexp_gas(4)[OF Some(9)] ffold_init_ad_same[of ct0 "(emptyEnv adv0 c0 (address ev) v0')" "(fmdom ct0)" "(ffold (init ct0) (emptyEnv adv0 c0 (address ev) v0) (fmdom ct0))"] unfolding emptyEnv_def by simp + moreover have "type (accounts (st\gas := g0'', accounts := acc0, stack := k\<^sub>l0, memory := m\<^sub>l0\) ad) = Some (Contract cname)" using transfer_type_same[OF Some(10)] a00 by simp + moreover have "iv (storage (st\gas := g0'', accounts := acc0, stack := k\<^sub>l0, memory := m\<^sub>l0\) ad) + (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts (st\gas := g0'', accounts := acc0, stack := k\<^sub>l0, memory := m\<^sub>l0\) ad)))" + proof - + from Some(5) have "c0 |\| fmdom ep" by (rule fmdomI) + with a0 a1 Some `adv0 \ ad` have "iv (storage st ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st ad)) - ReadL\<^sub>i\<^sub>n\<^sub>t v0')" by simp + moreover have "ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc0 ad)) = ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st ad)) - ReadL\<^sub>i\<^sub>n\<^sub>t v0'" using transfer_sub[OF Some(10)] a0 `adv0 \ ad` by simp + ultimately show ?thesis by simp + qed + moreover have "\st'::State. gas st' < gas (st\gas := g0'', accounts := acc0, stack := k\<^sub>l0, memory := m\<^sub>l0\) \ + (\mid fp f ev. + members $$ mid = Some (Method (fp, True, f)) \ + address ev \ ad + \ (\adex cd st0 xe val g v t g' v' e\<^sub>l cd\<^sub>l k\<^sub>l' m\<^sub>l' g'' acc. + g'' \ gas st' \ + type (acc ad) = Some (Contract cname) \ + local.expr adex ev cd (st0\gas := gas st0 - costs (EXTERNAL adex mid xe val) ev cd st0\) (gas st0 - costs (EXTERNAL adex mid xe val) ev cd st0) = Normal ((KValue ad, Value TAddr), g) \ + local.expr val ev cd (st0\gas := g\) g = Normal ((KValue v, Value t), g') \ + convert t (TUInt 256) v = Some v' \ + local.load True fp xe (ffold (init members) (emptyEnv ad cname (address ev) v') (fmdom members)) emptyStore emptyStore emptyStore ev cd (st0\gas := g'\) g' = Normal ((e\<^sub>l, cd\<^sub>l, k\<^sub>l', m\<^sub>l'), g'') \ + transfer (address ev) ad v' (accounts (st0\gas := g''\)) = Some acc \ + iv (storage st0 ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc ad)) - ReadL\<^sub>i\<^sub>n\<^sub>t v') + \ wpS (local.stmt f e\<^sub>l cd\<^sub>l) (\st. iv (storage st ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st ad)))) (\e. e = Gas \ e = Err) (st0\gas := g'', accounts := acc, stack := k\<^sub>l', memory := m\<^sub>l'\)))" (is "\st'. ?left st' \ ?right st'") + proof (rule allI,rule impI) + fix st'::State + assume "gas st' < gas (st\gas := g0'', accounts := acc0, stack := k\<^sub>l0, memory := m\<^sub>l0\)" + then have "gas st' < gas st" using msel_ssel_expr_load_rexp_gas(4)[OF Some(9)] msel_ssel_expr_load_rexp_gas(3)[OF Some(2)] msel_ssel_expr_load_rexp_gas(3)[OF Some(6)] by auto + then show "?right st'" using assm[OF all_gas_le[OF `gas st' < gas st` "1.IH"], THEN conjunct1] unfolding Qe_def by simp + qed + moreover have "\st'::State. gas st' < gas (st\gas := g0'', accounts := acc0, stack := k\<^sub>l0, memory := m\<^sub>l0\) \ + (\ev. address ev \ ad + \ (\ex cd st0 adex cc v t g g' v' acc. + g' \ gas st' \ + type (acc ad) = Some (Contract cname) \ + expr adex ev cd (st0\gas := gas st0 - cc\) (gas st0 - cc) = Normal ((KValue ad, Value TAddr), g) \ + expr ex ev cd (st0\gas := g\) g = Normal ((KValue v, Value t), g') \ + convert t (TUInt 256) v = Some v' \ + transfer (address ev) ad v' (accounts st0) = Some acc \ + iv (storage st0 ad) (\bal (acc ad)\ - \v'\) \ + wpS (local.stmt fb (ffold (init members) (emptyEnv ad cname (address ev) v') (fmdom members)) emptyStore) + (\st. iv (storage st ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st ad)))) (\e. e = Gas \ e = Err) + (st0\gas := g', accounts := acc, stack := emptyStore, memory := emptyStore\)))" (is "\st'. ?left st' \ ?right st'") + proof (rule allI,rule impI) + fix st'::State + assume l0: "gas st' < gas (st\gas := g0'', accounts := acc0, stack := k\<^sub>l0, memory := m\<^sub>l0\)" + then have "gas st' < gas st" using msel_ssel_expr_load_rexp_gas(4)[OF Some(9)] msel_ssel_expr_load_rexp_gas(3)[OF Some(2)] msel_ssel_expr_load_rexp_gas(3)[OF Some(6)] by auto + then show "?right st'" using assm[OF all_gas_le[OF `gas st' < gas st` "1.IH"], THEN conjunct2, THEN conjunct2, THEN conjunct2] unfolding Qfe_def by simp + qed + ultimately show ?thesis using safeStore[of e\<^sub>l0 ad "st\gas := g0'', accounts := acc0, stack := k\<^sub>l0, memory := m\<^sub>l0\" iv f0 cd\<^sub>l0 st0''] Some unfolding Qe_def Qfe_def by blast + qed + moreover have "gas st0'' \ gas st" using msel_ssel_expr_load_rexp_gas(4)[OF Some(9),THEN conjunct1] msel_ssel_expr_load_rexp_gas(3)[OF Some(2)] msel_ssel_expr_load_rexp_gas(3)[OF Some(6)] stmt_gas[OF Some(11)] by simp + ultimately show ?thesis by simp + next + case (None adv0 c0 g0 ct0 cn0 fb0' v0 t0 g0' v0' acc0 st0'') + moreover have "iv (storage s'''''' ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts s'''''' ad)))" + proof - + from None have "adv0 \ ad" using a0 by auto + then have "address (ffold (init ct0) (emptyEnv adv0 c0 (address ev) v0') (fmdom ct0)) \ ad" using ffold_init_ad_same[where ?e="\address = adv0, contract = c0, sender = address ev, svalue = v0', denvalue = {$$}\" and e'="ffold (init ct0) (emptyEnv adv0 c0 (address ev) v0') (fmdom ct0)"] unfolding emptyEnv_def by simp + moreover have "type (accounts (st\gas := g0', accounts := acc0, stack := emptyStore, memory := emptyStore\) ad) = Some (Contract cname)" using transfer_type_same[OF None(9)] a00 by simp + moreover have "iv (storage (st\gas := g0', accounts := acc0, stack := emptyStore, memory := emptyStore\) ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts (st\gas := g0', accounts := acc0, stack := emptyStore, memory := emptyStore\) ad)))" + proof - + from None(5) have "c0 |\| fmdom ep" by (rule fmdomI) + with a0 a1 None `adv0 \ ad` have "iv (storage st ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st ad)) - ReadL\<^sub>i\<^sub>n\<^sub>t v0')" by simp + moreover have "ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc0 ad)) = ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st ad)) - ReadL\<^sub>i\<^sub>n\<^sub>t v0'" using transfer_sub[OF None(9)] a0 `adv0 \ ad` by simp + ultimately show ?thesis by simp + qed + moreover have "\st'::State. gas st' < gas (st\gas := g0', accounts := acc0, stack := emptyStore, memory := emptyStore\) \ + (\mid fp f ev. + members $$ mid = Some (Method (fp, True, f)) \ + address ev \ ad + \ (\adex cd st0 xe val g v t g' v' e\<^sub>l cd\<^sub>l k\<^sub>l' m\<^sub>l' g'' acc. + g'' \ gas st' \ + type (acc ad) = Some (Contract cname) \ + local.expr adex ev cd (st0\gas := gas st0 - costs (EXTERNAL adex mid xe val) ev cd st0\) (gas st0 - costs (EXTERNAL adex mid xe val) ev cd st0) = Normal ((KValue ad, Value TAddr), g) \ + local.expr val ev cd (st0\gas := g\) g = Normal ((KValue v, Value t), g') \ + convert t (TUInt 256) v = Some v' \ + local.load True fp xe (ffold (init members) (emptyEnv ad cname (address ev) v') (fmdom members)) emptyStore emptyStore emptyStore ev cd (st0\gas := g'\) g' = Normal ((e\<^sub>l, cd\<^sub>l, k\<^sub>l', m\<^sub>l'), g'') \ + transfer (address ev) ad v' (accounts (st0\gas := g''\)) = Some acc \ + iv (storage st0 ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc ad)) - ReadL\<^sub>i\<^sub>n\<^sub>t v') + \ wpS (local.stmt f e\<^sub>l cd\<^sub>l) (\st. iv (storage st ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st ad)))) (\e. e = Gas \ e = Err) (st0\gas := g'', accounts := acc, stack := k\<^sub>l', memory := m\<^sub>l'\)))" (is "\st'. ?left st' \ ?right st'") + proof (rule allI,rule impI) + fix st'::State + assume "gas st' < gas (st\gas := g0', accounts := acc0, stack := emptyStore, memory := emptyStore\)" + then have "gas st' < gas st" using msel_ssel_expr_load_rexp_gas(3)[OF None(2)] msel_ssel_expr_load_rexp_gas(3)[OF None(6)] by auto + then show "?right st'" using assm[OF all_gas_le[OF `gas st' < gas st` "1.IH"], THEN conjunct1] unfolding Qe_def by simp + qed + moreover have "\st'::State. gas st' < gas (st\gas := g0', accounts := acc0, stack := emptyStore, memory := emptyStore\) \ + (\ev. address ev \ ad + \ (\ex cd st0 adex cc v t g g' v' acc. + g' \ gas st' \ + type (acc ad) = Some (Contract cname) \ + expr adex ev cd (st0\gas := gas st0 - cc\) (gas st0 - cc) = Normal ((KValue ad, Value TAddr), g) \ + expr ex ev cd (st0\gas := g\) g = Normal ((KValue v, Value t), g') \ + convert t (TUInt 256) v = Some v' \ + transfer (address ev) ad v' (accounts st0) = Some acc \ + iv (storage st0 ad) (\bal (acc ad)\ - \v'\) \ + wpS (local.stmt fb (ffold (init members) (emptyEnv ad cname (address ev) v') (fmdom members)) emptyStore) + (\st. iv (storage st ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st ad)))) (\e. e = Gas \ e = Err) + (st0\gas := g', accounts := acc, stack := emptyStore, memory := emptyStore\)))" (is "\st'. ?left st' \ ?right st'") + proof (rule allI,rule impI) + fix st'::State + assume l0: "gas st' < gas (st\gas := g0', accounts := acc0, stack := emptyStore, memory := emptyStore\)" + then have "gas st' < gas st" using msel_ssel_expr_load_rexp_gas(3)[OF None(2)] msel_ssel_expr_load_rexp_gas(3)[OF None(6)] by auto + then show "?right st'" using assm[OF all_gas_le[OF `gas st' < gas st` "1.IH"], THEN conjunct2, THEN conjunct2, THEN conjunct2] unfolding Qfe_def by simp + qed + ultimately have "iv (storage st0'' ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st0'' ad)))" using safeStore[of "ffold (init ct0) (emptyEnv adv0 c0 (address ev) v0') (fmdom ct0)" ad "st\gas := g0', accounts := acc0, stack := emptyStore, memory := emptyStore\" iv fb0' emptyStore "st0''"] None unfolding Qe_def Qfe_def by blast + then show ?thesis using None(11) by simp + qed + moreover have "gas st0'' \ gas st" using msel_ssel_expr_load_rexp_gas(3)[OF None(2)] msel_ssel_expr_load_rexp_gas(3)[OF None(6)] stmt_gas[OF None(10)] by simp + ultimately show ?thesis by simp + qed + next + fix e + assume "local.stmt (EXTERNAL ad' i xe val) ev cd st = Exception e" + show "e = Gas \ e = Err" using Ex.nchotomy by simp + qed + next + fix ev ex ad' cd + assume a00: "type (accounts st ad) = Some (Contract cname)" + and a0: "address ev = ad" + and a1: "\adv g. local.expr ad' ev cd (st\gas := gas st - costs (TRANSFER ad' ex) ev cd st\) (gas st - costs (TRANSFER ad' ex) ev cd st) = Normal ((KValue adv, Value TAddr), g) \ adv \ ad" + and a2: "\adv g v t g' v'. + local.expr ad' ev cd (st\gas := gas st - costs (TRANSFER ad' ex) ev cd st\) (gas st - costs (TRANSFER ad' ex) ev cd st) = Normal ((KValue adv, Value TAddr), g) \ + adv \ ad \ + local.expr ex ev cd (st\gas := g\) g = Normal ((KValue v, Value t), g') \ + convert t (TUInt 256) v = Some v' \ + iv (storage st ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st ad)) - ReadL\<^sub>i\<^sub>n\<^sub>t v')" + show "wpS (local.stmt (TRANSFER ad' ex) ev cd) (\st. iv (storage st ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st ad)))) (\e. e = Gas \ e = Err) st" + unfolding wpS_def wp_def + proof (split result.split; split prod.split; rule conjI; (rule allI)+; (rule impI)+) + fix x1 x1a s'''''' + assume "x1 = (x1a, s'''''')" and "local.stmt (TRANSFER ad' ex) ev cd st = Normal x1" + then have 2: "local.stmt (TRANSFER ad' ex) ev cd st = Normal (x1a, s'''''')" by simp + then show "gas s'''''' \ gas st \ iv (storage s'''''' ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts s'''''' ad)))" + proof (cases rule: transfer) + case (Contract v0 t0 g0 adv0 c0 g0' v0' acc0 ct0 cn0 f0 st0'') + moreover have "iv (storage s'''''' ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts s'''''' ad)))" + proof - + from Contract have "adv0 \ ad" using a1 by auto + then have "address (ffold (init ct0) (emptyEnv adv0 c0 (address ev) v0') (fmdom ct0)) \ ad" using a0 ffold_init_ad_same[where ?e'="ffold (init ct0) (emptyEnv adv0 c0 (address ev) v0') (fmdom ct0)"] unfolding emptyEnv_def by simp + moreover have "type (accounts (st\gas := g0', accounts := acc0, stack := emptyStore, memory := emptyStore\) ad) = Some (Contract cname)" using transfer_type_same[OF Contract(7)] a00 by simp + moreover have "iv (storage (st\gas := g0', accounts := acc0, stack := emptyStore, memory := emptyStore\) ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts (st\gas := g0', accounts := acc0, stack := emptyStore, memory := emptyStore\) ad)))" + proof - + from a0 a2 Contract `adv0 \ ad` have "iv (storage st ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st ad)) - ReadL\<^sub>i\<^sub>n\<^sub>t v0')" by blast + moreover have "ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc0 ad)) = ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st ad)) - ReadL\<^sub>i\<^sub>n\<^sub>t v0'" using transfer_sub[OF Contract(7)] a0 `adv0 \ ad` by simp + ultimately show ?thesis by simp + qed + moreover have "\st'::State. gas st' < gas (st\gas := g0', accounts := acc0, stack := emptyStore, memory := emptyStore\) \ Qe ad iv st'" + proof (rule allI,rule impI) + fix st'::State + assume "gas st' < gas (st\gas := g0', accounts := acc0, stack := emptyStore, memory := emptyStore\)" + then have "gas st' < gas st" using msel_ssel_expr_load_rexp_gas(3)[OF Contract(2)] msel_ssel_expr_load_rexp_gas(3)[OF Contract(3)] by auto + then show "Qe ad iv st'" using assm[OF all_gas_le[OF `gas st' < gas st` "1.IH"], THEN conjunct1] by simp + qed + moreover have "\st'::State. gas st' < gas (st\gas := g0', accounts := acc0, stack := emptyStore, memory := emptyStore\) \ + (\ev. address ev \ ad + \ (\ex cd st0 adex cc v t g g' v' acc. + g' \ gas st' \ + type (acc ad) = Some (Contract cname) \ + expr adex ev cd (st0\gas := gas st0 - cc\) (gas st0 - cc) = Normal ((KValue ad, Value TAddr), g) \ + expr ex ev cd (st0\gas := g\) g = Normal ((KValue v, Value t), g') \ + convert t (TUInt 256) v = Some v' \ + transfer (address ev) ad v' (accounts st0) = Some acc \ + iv (storage st0 ad) (\bal (acc ad)\ - \v'\) \ + wpS (local.stmt fb (ffold (init members) (emptyEnv ad cname (address ev) v') (fmdom members)) emptyStore) + (\st. iv (storage st ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st ad)))) (\e. e = Gas \ e = Err) + (st0\gas := g', accounts := acc, stack := emptyStore, memory := emptyStore\)))" (is "\st'. ?left st' \ ?right st'") + proof (rule allI,rule impI) + fix st'::State + assume l0: "gas st' < gas (st\gas := g0', accounts := acc0, stack := emptyStore, memory := emptyStore\)" + then have "gas st' < gas st" using msel_ssel_expr_load_rexp_gas(3)[OF Contract(2)] msel_ssel_expr_load_rexp_gas(3)[OF Contract(3)] by auto + then show "?right st'" using assm[OF all_gas_le[OF `gas st' < gas st` "1.IH"], THEN conjunct2, THEN conjunct2, THEN conjunct2] unfolding Qfe_def by simp + qed + ultimately have "iv (storage st0'' ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st0'' ad)))" using safeStore[of "ffold (init ct0) (emptyEnv adv0 c0 (address ev) v0') (fmdom ct0)" ad "st\gas := g0', accounts := acc0, stack := emptyStore, memory := emptyStore\" iv f0 emptyStore "st0''"] Contract unfolding Qe_def Qfe_def by blast + then show ?thesis using Contract(9) by simp + qed + moreover have "gas st0'' \ gas st" using msel_ssel_expr_load_rexp_gas(3)[OF Contract(2)] msel_ssel_expr_load_rexp_gas(3)[OF Contract(3)] stmt_gas[OF Contract(8)] by simp + ultimately show ?thesis by simp + next + case (EOA v0 t0 g0 adv0 g0' v0' acc0) + moreover have "iv (storage (st\gas := g0', accounts := acc0\) ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts (st\gas := g0', accounts := acc0\) ad)))" + proof - + from EOA have "adv0 \ ad" using a1 by auto + with a0 a2 EOA have "iv (storage st ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st ad)) - ReadL\<^sub>i\<^sub>n\<^sub>t v0')" by blast + moreover have "ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc0 ad)) = ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st ad)) - ReadL\<^sub>i\<^sub>n\<^sub>t v0'" using transfer_sub[OF EOA(6)] a0 `adv0 \ ad` by simp + ultimately show ?thesis by simp + qed + moreover have "g0' \ gas st" using msel_ssel_expr_load_rexp_gas(3)[OF EOA(2)] msel_ssel_expr_load_rexp_gas(3)[OF EOA(3)] by simp + ultimately show ?thesis by simp + qed + next + fix e + assume "local.stmt (TRANSFER ad' ex) ev cd st = Exception e" + show "e = Gas \ e = Err" using Ex.nchotomy by simp + qed + next + fix ev i xe cd fp + assume a0: "type (accounts st ad) = Some (Contract cname)" + and ad_ev: "address ev = ad" + and a1: "contract ev = cname" + and pre: "\fp e\<^sub>l cd\<^sub>l k\<^sub>l m\<^sub>l g. + local.load False fp xe (ffold (init members) (emptyEnv ad (contract ev) (sender ev) (svalue ev)) (fmdom members)) emptyStore emptyStore (memory st) ev cd (st\gas := gas st - costs (INVOKE i xe) ev cd st\) (gas st - costs (INVOKE i xe) ev cd st) = + Normal ((e\<^sub>l, cd\<^sub>l, k\<^sub>l, m\<^sub>l), g) \ + pre i (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st ad)), storage st ad, e\<^sub>l, cd\<^sub>l, k\<^sub>l, m\<^sub>l)" + show "wpS (local.stmt (INVOKE i xe) ev cd) (\st. post i (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st ad)), storage st ad)) (\e. e = Gas \ e = Err) st" + unfolding wpS_def wp_def + proof (split result.split; split prod.split; rule conjI; (rule allI)+; (rule impI)+) + fix x1 x1a st' + assume "x1 = (x1a, st')" + and *: "local.stmt (INVOKE i xe) ev cd st = Normal x1" + then have "local.stmt (INVOKE i xe) ev cd st = Normal (x1a, st')" by simp + then show "gas st' \ gas st \ post i (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st' ad)), storage st' ad)" + proof (cases rule: invoke) + case (1 ct fb fp f e\<^sub>l cd\<^sub>l k\<^sub>l m\<^sub>l g st'') + have "post i (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st' ad)), storage st' ad)" + proof - + from * have "gas st > costs (INVOKE i xe) ev cd st" by (simp add:stmt.simps split:if_split_asm) + then have "gas (st\gas := gas st - costs (INVOKE i xe) ev cd st\) < gas st" using invoke_not_zero[of i xe ev cd st] by simp + + from a1 have "ct = members" using 1(2) C1 by simp + then have **: "local.load False fp xe (ffold (init members) (emptyEnv ad (contract ev) (sender ev) (svalue ev)) (fmdom members)) emptyStore + emptyStore (memory st) ev cd (st\gas := gas st - costs (INVOKE i xe) ev cd st\) + (gas st - costs (INVOKE i xe) ev cd st) = + Normal ((e\<^sub>l, cd\<^sub>l, k\<^sub>l, m\<^sub>l), g)" using 1(4) ad_ev by simp + moreover from 1(2,3) have ***: "members $$ i = Some (Method (fp, False, f))" using ad_ev `ct = members` by simp + ultimately have "pre i (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st ad)), storage st ad, e\<^sub>l, cd\<^sub>l, k\<^sub>l, m\<^sub>l)" using pre by blast + moreover have "g \ gas (st\gas := gas st - costs (INVOKE i xe) ev cd st\)" using msel_ssel_expr_load_rexp_gas(4)[OF 1(4),THEN conjunct1] by simp + ultimately have "wpS (local.stmt f e\<^sub>l cd\<^sub>l) (\st. post i (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st ad)), storage st ad)) (\e. e = Gas \ e = Err) + (st\gas := g, stack := k\<^sub>l, memory := m\<^sub>l\)" using assm[OF all_gas_le[OF `gas (st\gas := gas st - costs (INVOKE i xe) ev cd st\) < gas st` "1.IH"], THEN conjunct2, THEN conjunct1] ** *** ad_ev a1 unfolding Qi_def by simp + then show ?thesis unfolding wpS_def wp_def using 1(5,6) by simp + qed + moreover have "gas st' \ gas st" using msel_ssel_expr_load_rexp_gas(4)[OF 1(4),THEN conjunct1] stmt_gas[OF 1(5)] 1(6) by simp + ultimately show ?thesis by simp + qed + next + fix e + assume "local.stmt (INVOKE i xe) ev cd st = Exception e" + show "e = Gas \ e = Err" using Ex.nchotomy by simp + qed + next + fix ev ex ad' cd + assume a0: "type (accounts st ad) = Some (Contract cname)" + and ad_ev: "address ev = ad" + and a1: "\adv g. + local.expr ad' ev cd (st\gas := gas st - costs (TRANSFER ad' ex) ev cd st\) + (gas st - costs (TRANSFER ad' ex) ev cd st) = Normal ((KValue adv, Value TAddr), g) \ adv = ad" + and a2: "\g v t g'. + local.expr ad' ev cd (st\gas := gas st - costs (TRANSFER ad' ex) ev cd st\) + (gas st - costs (TRANSFER ad' ex) ev cd st) = + Normal ((KValue ad, Value TAddr), g) \ + local.expr ex ev cd (st\gas := g\) g = Normal ((KValue v, Value t), g') \ + pref (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st ad)), storage st ad)" + show "wpS (local.stmt (TRANSFER ad' ex) ev cd) (\st. postf (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st ad)), storage st ad)) (\e. e = Gas \ e = Err) st" + unfolding wpS_def wp_def + proof (split result.split; split prod.split; rule conjI; (rule allI)+; (rule impI)+) + fix x1 x1a st' + assume "x1 = (x1a, st')" and "local.stmt (TRANSFER ad' ex) ev cd st = Normal x1" + then have 2: "local.stmt (TRANSFER ad' ex) ev cd st = Normal (x1a, st')" by simp + then show "gas st' \ gas st \ postf (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st' ad)), storage st' ad)" + proof (cases rule: transfer) + case (Contract v t g adv c g' v' acc ct cn f st'') + moreover from Contract have "adv = ad" using a1 by auto + ultimately have "pref (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st ad)), storage st ad)" using ad_ev a2 by auto + moreover have "ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st ad)) = ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc ad))" using transfer_same[OF Contract(7)] `adv = ad` ad_ev by simp + ultimately have "pref (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc ad)), storage st ad)" by simp + moreover from a0 have "c = cname" using Contract(5) `adv = ad` by simp + then have "ct = members" and "f = fb" using C1 Contract(6) by simp+ + moreover have "gas st \ costs (TRANSFER ad' ex) ev cd st" using 2 by (simp add:stmt.simps split:if_split_asm) + then have "gas (st\gas := gas st - costs (TRANSFER ad' ex) ev cd st\) < gas st" using transfer_not_zero[of ad' ex ev cd st] by simp + moreover have "g' \ gas (st\gas := gas st - costs (TRANSFER ad' ex) ev cd st\)" using msel_ssel_expr_load_rexp_gas(3)[OF Contract(2)] msel_ssel_expr_load_rexp_gas(3)[OF Contract(3)] by simp + ultimately have "wpS (local.stmt fb (ffold (init members) (emptyEnv ad c (address ev) v') (fmdom members)) emptyStore) + (\st. postf (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st ad)), storage st ad)) (\e. e = Gas \ e = Err) + (st\gas := g', accounts := acc, stack := emptyStore, memory := emptyStore\)" using assm[OF all_gas_le[OF `gas (st\gas := gas st - costs (TRANSFER ad' ex) ev cd st\) < gas st` "1.IH"], THEN conjunct2, THEN conjunct2, THEN conjunct1] ad_ev Contract `adv = ad` `c = cname` unfolding Qfi_def by blast + with `ct = members` `f=fb` have "gas st' \ gas (st\gas := g', accounts := acc, stack := emptyStore, memory := emptyStore\) \ postf (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st' ad)), storage st' ad)" unfolding wpS_def wp_def using Contract(8,9) `adv = ad` by simp + moreover from this have "gas st' \ gas st" using `g' \ gas (st\gas := gas st - costs (TRANSFER ad' ex) ev cd st\)` by auto + ultimately show ?thesis by simp + next + case (EOA v t g adv g' acc) + then show ?thesis using a0 a1 by simp + qed + next + fix e + assume "local.stmt (TRANSFER ad' ex) ev cd st = Exception e" + show "e = Gas \ e = Err" using Ex.nchotomy by simp + qed + qed +qed + +text \ + Refined versions of @{thm[source] wp_external_invoke_transfer}. +\ + +lemma wp_transfer_ext[rule_format]: + assumes "\st::State. \\st'::State. gas st' \ gas st \ type (accounts st' ad) = Some (Contract cname) \ Pe ad iv st' \ Pi ad pre post st' \ Pfi ad pref postf st' \ Pfe ad iv st'\ + \ Qe ad iv st \ Qi ad pre post st \ Qfi ad pref postf st \ Qfe ad iv st" + and "type (accounts st ad) = Some (Contract cname)" + shows "(\ev ex ad' cd. + address ev = ad \ + (\adv g. + expr ad' ev cd (st\gas := gas st - costs (TRANSFER ad' ex) ev cd st\) (gas st - costs (TRANSFER ad' ex) ev cd st) = Normal ((KValue adv, Value TAddr), g) + \ adv \ ad) \ + (\adv g v t g' v'. + expr ad' ev cd (st\gas := gas st - costs (TRANSFER ad' ex) ev cd st\) (gas st - costs (TRANSFER ad' ex) ev cd st) = Normal ((KValue adv, Value TAddr), g) \ + adv \ ad \ + expr ex ev cd (st\gas := g\) g = Normal ((KValue v, Value t), g') \ + convert t (TUInt 256) v = Some v' + \ iv (storage st ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st ad)) - ReadL\<^sub>i\<^sub>n\<^sub>t v')) + \ wpS (\s. stmt (TRANSFER ad' ex) ev cd s) (\st. iv (storage st ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st ad)))) (\e. e = Gas \ e = Err) st)" +proof - + from wp_external_invoke_transfer have "Pfe ad iv st" using assms by blast + then show ?thesis using Pfe_def by simp +qed + +lemma wp_external[rule_format]: + assumes "\st::State. \\st'::State. gas st' \ gas st \ type (accounts st' ad) = Some (Contract cname)\ Pe ad iv st' \ Pi ad pre post st' \ Pfi ad pref postf st' \ Pfe ad iv st'\ + \ Qe ad iv st \ Qi ad pre post st \ Qfi ad pref postf st \ Qfe ad iv st" + and "type (accounts st ad) = Some (Contract cname)" + shows "(\ev ad' i xe val cd. + address ev = ad \ + (\adv c g v t g' v'. + expr ad' ev cd (st\gas := gas st - costs (EXTERNAL ad' i xe val) ev cd st\) (gas st - costs (EXTERNAL ad' i xe val) ev cd st) = Normal ((KValue adv, Value TAddr), g) \ + adv \ ad \ + type (accounts st adv) = Some (Contract c) \ + c |\| fmdom ep \ + expr val ev cd (st\gas := g\) g = Normal ((KValue v, Value t), g') \ + convert t (TUInt 256) v = Some v' + \ iv (storage st ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st ad)) - ReadL\<^sub>i\<^sub>n\<^sub>t v')) + \ wpS (\s. stmt (EXTERNAL ad' i xe val) ev cd s) (\st. iv (storage st ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st ad)))) (\e. e = Gas \ e = Err) st)" +proof - + from wp_external_invoke_transfer have "Pe ad iv st" using assms by blast + then show ?thesis using Pe_def by simp +qed + +lemma wp_invoke[rule_format]: + assumes "\st::State. \\st'::State. gas st' \ gas st \ type (accounts st' ad) = Some (Contract cname) \ Pe ad iv st' \ Pi ad pre post st' \ Pfi ad pref postf st' \ Pfe ad iv st'\ + \ Qe ad iv st \ Qi ad pre post st \ Qfi ad pref postf st \ Qfe ad iv st" + and "type (accounts st ad) = Some (Contract cname)" + shows "(\ev i xe cd. + address ev = ad \ + contract ev = cname \ + (\fp e\<^sub>l cd\<^sub>l k\<^sub>l m\<^sub>l g. + load False fp xe (ffold (init members) (emptyEnv ad (contract ev) (sender ev) (svalue ev)) (fmdom members)) emptyStore emptyStore (memory st) ev cd (st\gas := gas st - costs (INVOKE i xe) ev cd st\) (gas st - costs (INVOKE i xe) ev cd st) = Normal ((e\<^sub>l, cd\<^sub>l, k\<^sub>l, m\<^sub>l), g) + \ pre i (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st ad)), storage st ad, e\<^sub>l, cd\<^sub>l, k\<^sub>l, m\<^sub>l)) + \ wpS (\s. stmt (INVOKE i xe) ev cd s) (\st. post i (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st ad)), storage st ad)) (\e. e = Gas \ e = Err) st)" +proof - + from wp_external_invoke_transfer have "Pi ad pre post st" using assms by blast + then show ?thesis using Pi_def by simp +qed + +lemma wp_transfer_int[rule_format]: + assumes "\st::State. \\st'::State. gas st' \ gas st \ type (accounts st' ad) = Some (Contract cname) \ Pe ad iv st' \ Pi ad pre post st' \ Pfi ad pref postf st' \ Pfe ad iv st'\ + \ Qe ad iv st \ Qi ad pre post st \ Qfi ad pref postf st \ Qfe ad iv st" + and "type (accounts st ad) = Some (Contract cname)" + shows "(\ev ex ad' cd. + address ev = ad \ + (\adv g. + expr ad' ev cd (st\gas := gas st - costs (TRANSFER ad' ex) ev cd st\) (gas st - costs (TRANSFER ad' ex) ev cd st) = Normal ((KValue adv, Value TAddr), g) + \ adv = ad) \ + (\g v t g'. + expr ad' ev cd (st\gas := gas st - costs (TRANSFER ad' ex) ev cd st\) (gas st - costs (TRANSFER ad' ex) ev cd st) = Normal ((KValue ad, Value TAddr), g) \ + expr ex ev cd (st\gas := g\) g = Normal ((KValue v, Value t), g') + \ pref (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st ad)), storage st ad)) + \ wpS (\s. stmt (TRANSFER ad' ex) ev cd s) (\st. postf (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st ad)), storage st ad)) (\e. e = Gas \ e = Err) st)" +proof - + from wp_external_invoke_transfer have "Pfi ad pref postf st" using assms by blast + then show ?thesis using Pfi_def by simp +qed + +definition external :: "Address \ ((String.literal, String.literal) fmap \ int \ bool) \ bool" + where "external ad iv \ (\acc g'' m\<^sub>l k\<^sub>l cd\<^sub>l e\<^sub>l g' t v v' g val xe cd adex ev f fp i st. + type (accounts st ad) = Some (Contract cname) \ + members $$ i = Some (Method (fp, True, f)) \ + address ev \ ad \ + expr adex ev cd (st\gas := gas st - costs (EXTERNAL adex i xe val) ev cd st\) (gas st - costs (EXTERNAL adex i xe val) ev cd st) = Normal ((KValue ad, Value TAddr), g) \ + expr val ev cd (st\gas := g\) g = Normal ((KValue v, Value t), g') \ + convert t (TUInt 256) v = Some v' \ + load True fp xe (ffold (init members) (emptyEnv ad cname (address ev) v') (fmdom members)) emptyStore emptyStore emptyStore ev cd (st\gas := g'\) g' = Normal ((e\<^sub>l, cd\<^sub>l, k\<^sub>l, m\<^sub>l), g'') \ + transfer (address ev) ad v' (accounts (st\gas := g''\)) = Some acc \ + iv (storage st ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc ad)) - ReadL\<^sub>i\<^sub>n\<^sub>t v') + \ (wpS (\s. stmt f e\<^sub>l cd\<^sub>l s) (\st. iv (storage st ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st ad)))) (\e. e = Gas \ e=Err) (st\gas := g'', accounts := acc, stack:=k\<^sub>l, memory:=m\<^sub>l\)))" + +definition fallback :: "Address \ ((String.literal, String.literal) fmap \ int \ bool) \ bool" + where "fallback ad iv \ (\acc g' t v v' g val cd adex ev st k. + type (accounts st ad) = Some (Contract cname) \ + address ev \ ad \ + local.expr adex ev cd (st\gas := gas st - k\) (gas st - k) = + Normal ((KValue ad, Value TAddr), g) \ + local.expr val ev cd (st\gas := g\) g = Normal ((KValue v, Value t), g') \ + convert t (TUInt 256) v = Some v' \ + transfer (address ev) ad v' (accounts st) = Some acc \ + iv (storage st ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc ad)) - ReadL\<^sub>i\<^sub>n\<^sub>t v') + \ wpS (local.stmt fb (ffold (init members) (emptyEnv ad cname (address ev) v') (fmdom members)) emptyStore) (\st. iv (storage st ad) \bal (accounts st ad)\) (\e. e = Gas \ e = Err) + (st\gas := g', accounts := acc, stack := emptyStore, memory := emptyStore\))" + +definition constructor :: "((String.literal, String.literal) fmap \ int \ bool) \ bool" + where "constructor iv \ (\acc g'' m\<^sub>l k\<^sub>l cd\<^sub>l e\<^sub>l g' t v xe i cd val st ev adv. + adv = hash (address ev) (ShowL\<^sub>n\<^sub>a\<^sub>t (contracts (accounts st (address ev)))) \ + type (accounts st adv) = None \ + expr val ev cd (st\gas := gas st - costs (NEW i xe val) ev cd st\) (gas st - costs (NEW i xe val) ev cd st) = Normal ((KValue v, Value t), g') \ + load True (fst const) xe (ffold (init members) (emptyEnv adv cname (address ev) v) (fmdom members)) emptyStore emptyStore emptyStore ev cd (st\gas := g'\) g' = Normal ((e\<^sub>l, cd\<^sub>l, k\<^sub>l, m\<^sub>l), g'') \ + transfer (address ev) adv v (accounts (st\accounts := (accounts st)(adv := \bal = ShowL\<^sub>i\<^sub>n\<^sub>t 0, type = Some (Contract i), contracts = 0\)\)) = Some acc + \ wpS (local.stmt (snd const) e\<^sub>l cd\<^sub>l) (\st. iv (storage st adv) \bal (accounts st adv)\) (\e. e = Gas \ e = Err) + (st\gas := g'', storage:=(storage st)(adv := {$$}), accounts := acc, stack:=k\<^sub>l, memory:=m\<^sub>l\))" + +lemma sound: + fixes iv ad + assumes "\ad. external ad iv" + and "\ad. fallback ad iv" + and "constructor iv" + and "address ev \ ad" + and "type (accounts st ad) = Some (Contract cname) \ iv (storage st ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st ad)))" + shows "\(st'::State). stmt f ev cd st = Normal ((), st') \ type (accounts st' ad) = Some (Contract cname) + \ iv (storage st' ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st' ad)))" + using assms(4-) +proof (induction rule:stmt.induct) + case (1 ev cd st) + show ?case + proof elims + fix st' + assume *: "stmt SKIP ev cd st = Normal ((), st')" + and "type (accounts st' ad) = Some (Contract cname)" + then show "iv (storage st' ad) \bal (accounts st' ad)\" using 1 skip[OF *] by simp + qed +next + case (2 lv ex ev cd st) + show ?case + proof elims + fix st' + assume *: "stmt (ASSIGN lv ex) ev cd st = Normal ((), st')" + and "type (accounts st' ad) = Some (Contract cname)" + then show "iv (storage st' ad) \bal (accounts st' ad)\" using 2 by (cases rule: assign[OF *];simp) + qed +next + case (3 s1 s2 ev cd st) + show ?case + proof elims + fix st' + assume *: "stmt (COMP s1 s2) ev cd st = Normal ((), st')" + and **: "type (accounts st' ad) = Some (Contract cname)" + show "iv (storage st' ad) \bal (accounts st' ad)\" + proof (cases rule: comp[OF *]) + case (1 st'') + moreover from 3(4) have "type (accounts (st\gas := gas st - costs (COMP s1 s2) ev cd st\) ad) = Some (Contract cname) \ iv (storage (st\gas := gas st - costs (COMP s1 s2) ev cd st\) ad) \bal (accounts (st\gas := gas st - costs (COMP s1 s2) ev cd st\) ad)\" by auto + ultimately have "type (accounts st'' ad) = Some (Contract cname) \ iv (storage st'' ad) \bal (accounts st'' ad)\" using 3(1)[OF _ _ 3(3)] by fastforce + then show ?thesis using 3(2)[OF _ _ _ 3(3)] 1 ** by fastforce + qed + qed +next + case (4 ex s1 s2 ev cd st) + show ?case + proof elims + fix st' + assume a1: "local.stmt (ITE ex s1 s2) ev cd st = Normal ((), st')" + and a2: "type (accounts st' ad) = Some (Contract cname)" + show "iv (storage st' ad) \bal (accounts st' ad)\" + proof (cases rule:ite[OF a1]) + case (1 g) + have "\st'. local.stmt s1 ev cd (st\gas := g\) = Normal ((), st') \ + type (accounts st' ad) = Some (Contract cname) \ + iv (storage st' ad) \bal (accounts st' ad)\" + apply (rule 4(1)) using 1 4(3) 4(4) by auto + then show ?thesis using a2 1(3) by blast + next + case (2 g) + have "\st'. local.stmt s2 ev cd (st\gas := g\) = Normal ((), st') \ + type (accounts st' ad) = Some (Contract cname) \ + iv (storage st' ad) \bal (accounts st' ad)\" + apply (rule 4(2)) using 2 4(3) 4(4) true_neq_false[symmetric] by auto + then show ?thesis using a2 2(3) by blast + qed + qed +next + case (5 ex s0 ev cd st) + show ?case + proof elims + fix st' + assume a1: "local.stmt (WHILE ex s0) ev cd st = Normal ((), st')" + and a2: "type (accounts st' ad) = Some (Contract cname)" + show "iv (storage st' ad) \bal (accounts st' ad)\" + proof (cases rule:while[OF a1]) + case (1 g st'') + have "\st'. local.stmt s0 ev cd (st\gas := g\) = Normal ((), st') \ + type (accounts st' ad) = Some (Contract cname) \ + iv (storage st' ad) \bal (accounts st' ad)\" + apply (rule 5(1)) using 1 5(3) 5(4) by auto + then have *: "type (accounts st'' ad) = Some (Contract cname) \ + iv (storage st'' ad) \bal (accounts st'' ad)\" using 1(3) by simp + have "\st'. local.stmt (WHILE ex s0) ev cd st'' = Normal ((), st') \ + type (accounts st' ad) = Some (Contract cname) \ + iv (storage st' ad) \bal (accounts st' ad)\" + apply (rule 5(2)) using 1 5(3) 5(4) * by auto + then show ?thesis using a2 1(4) by blast + next + case (2 g) + then show ?thesis using a2 5(4) by simp + qed + qed +next + case (6 i xe ev cd st) + show ?case + proof elims + fix st' + assume a1: "local.stmt (INVOKE i xe) ev cd st = Normal ((), st')" + and a2: "type (accounts st' ad) = Some (Contract cname)" + show "iv (storage st' ad) \bal (accounts st' ad)\" + proof (cases rule:invoke[OF a1]) + case (1 ct fb fp f e\<^sub>l cd\<^sub>l k\<^sub>l m\<^sub>l g st'') + from 6(2) have "ad \ address e\<^sub>l" using msel_ssel_expr_load_rexp_gas(4)[OF 1(4)] ffold_init_ad by simp + have "\st'. local.stmt f e\<^sub>l cd\<^sub>l (st\gas := g, stack := k\<^sub>l, memory := m\<^sub>l\) = Normal ((), st') \ type (accounts st' ad) = Some (Contract cname) \ + iv (storage st' ad) \bal (accounts st' ad)\" apply (rule 6(1)) using 1 6(3) `ad \ address e\<^sub>l` by auto + then show ?thesis using a2 1(5,6) by auto + qed + qed +next + case (7 adex i xe val ev cd st) + show ?case + proof elims + fix st' + assume a1: "local.stmt (EXTERNAL adex i xe val) ev cd st = Normal ((), st')" + and a2: "type (accounts st' ad) = Some (Contract cname)" + show "iv (storage st' ad) \bal (accounts st' ad)\" + proof (cases rule:external[OF a1]) + case (1 adv c g ct cn fb' v t g' v' fp f e\<^sub>l cd\<^sub>l k\<^sub>l m\<^sub>l g'' acc st'') + then show ?thesis + proof (cases "adv = ad") + case True + then have "type (acc ad) = Some (Contract c)" using transfer_type_same[OF 1(10)] 1(4) by auto + then have "type (accounts st' ad) = Some (Contract c)" using atype_same[OF 1(11)] 1(12) by simp + then have "c = cname" using a2 by simp + then have "type (accounts st ad) = Some (Contract cname)" using 1(4) True by simp + moreover from `c = cname` have "ct = members" using 1 C1 by simp + moreover have "iv (storage st ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc ad)) - ReadL\<^sub>i\<^sub>n\<^sub>t v')" + proof - + have "iv (storage st ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st ad)))" using 7(4) `type (accounts st ad) = Some (Contract cname)` by simp + moreover have "ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc ad)) = ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st ad)) + ReadL\<^sub>i\<^sub>n\<^sub>t v'" using transfer_add[OF 1(10)] 7(3) True by simp + ultimately show ?thesis by simp + qed + ultimately have "wpS (local.stmt f e\<^sub>l cd\<^sub>l) (\st. iv (storage st ad) \bal (accounts st ad)\) (\e. e = Gas \ e = Err) + (st\gas := g'', accounts := acc, stack := k\<^sub>l, memory := m\<^sub>l\)" using assms(1) 1 True 7(3) `c = cname` unfolding external_def by auto + then show ?thesis unfolding wpS_def wp_def using 1(11,12) by simp + next + case False + then have *: "ad \ address e\<^sub>l" using msel_ssel_expr_load_rexp_gas(4)[OF 1(9)] ffold_init_ad by simp + moreover have **:"type (acc ad) = Some (Contract cname) \ iv (storage st ad) \bal (acc ad)\" + proof + assume "type (acc ad) = Some (Contract cname)" + then have "type (accounts st ad) = Some (Contract cname)" using transfer_type_same[OF 1(10)] by simp + then have "iv (storage st ad) \bal (accounts st ad)\" using 7(4) by simp + moreover have "bal (acc ad) = bal (accounts st ad)" using transfer_eq[OF 1(10)] 7(3) False by simp + ultimately show "iv (storage st ad) \bal (acc ad)\" by simp + qed + ultimately have "\st'. local.stmt f e\<^sub>l cd\<^sub>l (st\gas := g'', accounts := acc, stack := k\<^sub>l, memory := m\<^sub>l\) = Normal ((), st') \ + type (accounts st' ad) = Some (Contract cname) \ iv (storage st' ad) \bal (accounts st' ad)\" + using 7(1) 1 by auto + then show ?thesis using a2 1(11,12) by simp + qed + next + case (2 adv c g ct cn fb' v t g' v' acc st'') + then show ?thesis + proof (cases "adv = ad") + case True + then have "type (acc ad) = Some (Contract c)" using transfer_type_same[OF 2(9)] 2(4) by auto + then have "type (accounts st' ad) = Some (Contract c)" using atype_same[OF 2(10)] 2(11) by simp + then have "c = cname" using a2 by simp + then have "type (accounts st ad) = Some (Contract cname)" using 2(4) True by simp + moreover from `c = cname` have "ct = members" and "fb'=fb" using 2 C1 by simp+ + moreover have "iv (storage st ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc ad)) - ReadL\<^sub>i\<^sub>n\<^sub>t v')" + proof - + have "iv (storage st ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st ad)))" using 7(4) `type (accounts st ad) = Some (Contract cname)` by simp + moreover have "ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc ad)) = ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st ad)) + ReadL\<^sub>i\<^sub>n\<^sub>t v'" using transfer_add[OF 2(9)] 7(3) True by simp + ultimately show "iv (storage st ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc ad)) - ReadL\<^sub>i\<^sub>n\<^sub>t v')" by simp + qed + ultimately have "wpS (local.stmt fb' (ffold (init ct) (emptyEnv adv c (address ev) v') (fmdom ct)) emptyStore) (\st. iv (storage st ad) \bal (accounts st ad)\) (\e. e = Gas \ e = Err) + (st\gas := g', accounts := acc, stack := emptyStore, memory := emptyStore\)" using assms(2) 7(3) 2 True `c = cname` unfolding fallback_def by blast + then show ?thesis unfolding wpS_def wp_def using 2(10,11) by simp + next + case False + moreover have **:"type (acc ad) = Some (Contract cname) \ iv (storage st ad) \bal (acc ad)\" + proof + assume "type (acc ad) = Some (Contract cname)" + then have "type (accounts st ad) = Some (Contract cname)" using transfer_type_same[OF 2(9)] by simp + then have "iv (storage st ad) \bal (accounts st ad)\" using 7(4) by simp + moreover have "bal (acc ad) = bal (accounts st ad)" using transfer_eq[OF 2(9)] 7(3) False by simp + ultimately show "iv (storage st ad) \bal (acc ad)\" by simp + qed + ultimately have "\st'. local.stmt fb' (ffold (init ct) (emptyEnv adv c (address ev) v') (fmdom ct)) emptyStore (st\gas := g', accounts := acc, stack := emptyStore, memory := emptyStore\) = Normal ((), st') \ + type (accounts st' ad) = Some (Contract cname) \ iv (storage st' ad) \bal (accounts st' ad)\" + using 7(2) 2 by auto + then show ?thesis using a2 2(10,11) by simp + qed + qed + qed +next + case (8 ad' ex ev cd st) + show ?case + proof elims + fix st' + assume a1: "local.stmt (TRANSFER ad' ex) ev cd st = Normal ((), st')" + and a2: "type (accounts st' ad) = Some (Contract cname)" + show "iv (storage st' ad) \bal (accounts st' ad)\" + proof (cases rule:transfer[OF a1]) + case (1 v t g adv c g' v' acc ct cn f st'') + then show ?thesis + proof (cases "adv = ad") + case True + then have "type (acc ad) = Some (Contract c)" using transfer_type_same[OF 1(7)] 1(5) by auto + then have "type (accounts st' ad) = Some (Contract c)" using atype_same[OF 1(8)] 1(9) by simp + then have "c = cname" using a2 by simp + then have "type (accounts st ad) = Some (Contract cname)" using 1(5) True by simp + moreover from `c = cname` have "ct = members" and "f=fb" using 1 C1 by simp+ + moreover have "iv (storage st ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc ad)) - ReadL\<^sub>i\<^sub>n\<^sub>t v')" + proof - + have "iv (storage st ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st ad)))" using 8(3) `type (accounts st ad) = Some (Contract cname)` by simp + moreover have "ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc ad)) = ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st ad)) + ReadL\<^sub>i\<^sub>n\<^sub>t v'" using transfer_add[OF 1(7)] 8(2) True by simp + ultimately show "iv (storage st ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (acc ad)) - ReadL\<^sub>i\<^sub>n\<^sub>t v')" by simp + qed + ultimately have "wpS (local.stmt f (ffold (init ct) (emptyEnv adv c (address ev) v') (fmdom ct)) emptyStore) (\st. iv (storage st ad) \bal (accounts st ad)\) (\e. e = Gas \ e = Err) + (st\gas := g', accounts := acc, stack := emptyStore, memory := emptyStore\)" using assms(2) 8(2) 1 True `c = cname` unfolding fallback_def by blast + then show ?thesis unfolding wpS_def wp_def using 1(8,9) by simp + next + case False + moreover have "type (acc ad) = Some (Contract cname) \ iv (storage st ad) \bal (acc ad)\" + proof + assume "type (acc ad) = Some (Contract cname)" + then have "type (accounts st ad) = Some (Contract cname)" using transfer_type_same[OF 1(7)] by simp + then have "iv (storage st ad) \bal (accounts st ad)\" using 8(3) by simp + moreover have "bal (acc ad) = bal (accounts st ad)" using transfer_eq[OF 1(7)] 8(2) False by simp + ultimately show "iv (storage st ad) \bal (acc ad)\" by simp + qed + ultimately have "\st'. local.stmt f (ffold (init ct) (emptyEnv adv c (address ev) v') (fmdom ct)) emptyStore + (st\gas := g', accounts := acc, stack := emptyStore, memory := emptyStore\) = Normal ((), st') \ type (accounts st' ad) = Some (Contract cname) \ + iv (storage st' ad) \bal (accounts st' ad)\" using 8(1) 1 by auto + then show ?thesis using a2 1(8, 9) by simp + qed + next + case (2 v t g adv g' v' acc) + then show ?thesis + proof (cases "adv = ad") + case True + then show ?thesis using 2(5,7) a2 transfer_type_same[OF 2(6)] by simp + next + case False + then have "bal (acc ad) = bal (accounts st ad)" using transfer_eq[OF 2(6)] 8(2) by simp + moreover have "type (accounts st ad) = Some (Contract cname)" using transfer_type_same[OF 2(6)] 2(7) a2 by simp + then have "iv (storage st ad) \bal (accounts st ad)\" using 8(3) by simp + ultimately show ?thesis using 2(7) by simp + qed + qed + qed +next + case (9 id0 tp s e cd st) + show ?case + proof elims + fix st' + assume a1: "local.stmt (BLOCK ((id0, tp), None) s) e cd st = Normal ((), st')" + and a2: "type (accounts st' ad) = Some (Contract cname)" + show "iv (storage st' ad) \bal (accounts st' ad)\" + proof (cases rule:blockNone[OF a1]) + case (1 cd' mem' sck' e') + have "address e' = address e" using decl_env[OF 1(2)] by simp + have "\st'. local.stmt s e' cd' (st\gas := gas st - costs (BLOCK ((id0, tp), None) s) e cd st, stack := sck', + memory := mem'\) = Normal ((), st') \ + type (accounts st' ad) = Some (Contract cname) \ + iv (storage st' ad) \bal (accounts st' ad)\" + apply (rule 9(1)) using 1 9(2,3) `address e' = address e` by auto + then show ?thesis using a2 1(3) by blast + qed + qed +next + case (10 id0 tp ex' s e cd st) + show ?case + proof elims + fix st' + assume a1: "local.stmt (BLOCK ((id0, tp), Some ex') s) e cd st = Normal ((), st')" + and a2: "type (accounts st' ad) = Some (Contract cname)" + show "iv (storage st' ad) \bal (accounts st' ad)\" + proof (cases rule:blockSome[OF a1]) + case (1 v t g cd' mem' sck' e') + have "address e' = address e" using decl_env[OF 1(3)] by simp + have "\st'. local.stmt s e' cd' (st\gas := g, stack := sck', memory := mem'\) = Normal ((), st') \ + type (accounts st' ad) = Some (Contract cname) \ + iv (storage st' ad) \bal (accounts st' ad)\" + apply (rule 10(1)) using 1 10(2,3) `address e' = address e` by auto + then show ?thesis using a2 1(4) by blast + qed + qed +next + case (11 i xe val e cd st) + show ?case + proof elims + fix st' + assume a1: "local.stmt (NEW i xe val) e cd st = Normal ((), st')" + and a2: "type (accounts st' ad) = Some (Contract cname)" + show "iv (storage st' ad) \bal (accounts st' ad)\" + proof (cases rule:new[OF a1]) + case (1 v t g ct cn fb' e\<^sub>l cd\<^sub>l k\<^sub>l m\<^sub>l g' acc st'') + define adv where "adv = hash (address e) \contracts (accounts (st\gas := gas st - costs (NEW i xe val) e cd st\) (address e))\" + then have "address e\<^sub>l = adv" using msel_ssel_expr_load_rexp_gas(4)[OF 1(5)] by simp + then show ?thesis + proof (cases "adv = ad") + case True + then show ?thesis + proof (cases "i = cname") + case t0: True + then have "ct = members" and "cn = const" and "fb' = fb" using 1(4) C1 by simp+ + then have "wpS (local.stmt (snd const) e\<^sub>l cd\<^sub>l) (\st. iv (storage st adv) \bal (accounts st adv)\) (\e. e = Gas \ e = Err) + (st\gas := g', storage:=(storage st)(adv := {$$}), accounts := acc, stack:=k\<^sub>l, memory:=m\<^sub>l\)" + using assms(3) 11(3) 1 True adv_def t0 unfolding constructor_def by auto + then have "iv (storage st'' adv) \bal (accounts st'' adv)\" unfolding wpS_def wp_def using 1(7) `cn = const` adv_def by simp + then show ?thesis using 1(8) True by simp + next + case False + moreover have "type (accounts st' ad) = Some (Contract i)" + proof - + from `adv = ad` have "type (((accounts st) (adv := \bal = ShowL\<^sub>i\<^sub>n\<^sub>t 0, type = Some (Contract i), contracts = 0\)) ad) = Some (Contract i)" by simp + then have "type (acc ad) = Some (Contract i)" using transfer_type_same[OF 1(6)] adv_def by simp + then have "type (accounts st'' ad) = Some (Contract i)" using atype_same[OF 1(7)] adv_def by simp + then show ?thesis using 1(8) by simp + qed + ultimately show ?thesis using a2 by simp + qed + next + case False + moreover have *: "type (acc ad) = Some (Contract cname) \ iv (storage (st\storage:=(storage st)(adv := {$$}), accounts:=acc\) ad) \bal (acc ad)\" + proof + assume "type (acc ad) = Some (Contract cname)" + then have "type (((accounts st) (adv := \bal = ShowL\<^sub>i\<^sub>n\<^sub>t 0, type = Some (Contract i), contracts = 0\)) ad) = Some (Contract cname)" using transfer_type_same[OF 1(6)] adv_def by simp + then have "type ((accounts st) ad) = Some (Contract cname)" using False by simp + then have "iv (storage st ad) \bal (accounts st ad)\" using 11(3) by simp + then have "iv (storage (st\storage:=(storage st)(adv := {$$})\) ad) \bal (accounts st ad)\" using False by simp + then show "iv (storage (st\storage:=(storage st)(adv := {$$}), accounts:=acc\) ad) \bal (acc ad)\" using transfer_eq[OF 1(6)] adv_def 11(2) False by simp + qed + ultimately have "\st'. stmt (snd cn) e\<^sub>l cd\<^sub>l (st\gas := g', storage := (storage st) (adv := {$$}), accounts := acc, stack := k\<^sub>l, memory := m\<^sub>l\) = Normal ((), st') \ + type (accounts st' ad) = Some (Contract cname) \ iv (storage st' ad) \bal (accounts st' ad)\" + using 11(1)[where ?s'k4="st\gas := g', storage := (storage st) (adv := {$$}), accounts := acc, stack := k\<^sub>l, memory := m\<^sub>l\"] + 1 adv_def `address e\<^sub>l = adv` False * by auto + moreover have "type (accounts st'' ad) = Some (Contract cname)" using 1(8) a2 adv_def by auto + ultimately show ?thesis using a2 1(6,7,8) adv_def by simp + qed + qed + qed +qed + +lemma sound2: + fixes iv ad + assumes "\ad. external ad iv" + and "\ad. fallback ad iv" + and "constructor iv" + and "\ad. address ev \ ad \ type (accounts st ad) = Some (Contract cname) \ iv (storage st ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st ad)))" + shows "\(st'::State) ad. stmt f ev cd st = Normal ((), st') \ type (accounts st' ad) = Some (Contract cname) \ address ev \ ad + \ iv (storage st' ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st' ad)))" + using assms sound by blast +end + +context Calculus +begin + + named_theorems mcontract + named_theorems external + named_theorems internal + + section \Verification Condition Generator\ + + text \ + To use the verification condition generator first invoke the following rule on the original Hoare triple: + \ + method vcg_valid = + rule wpS_valid, + erule conjE, + simp + + method external uses cases = + unfold external_def, + elims, + (erule cases;simp) + + method fallback uses cases = + unfold fallback_def, + elims, + rule cases + + method constructor uses cases = + unfold constructor_def, + elims, + rule cases, + simp + + text \ + Then apply the correct rules from the following set of rules. + \ + + subsection \Skip\ + + method vcg_skip = + rule wp_Skip;(solve simp)? + + subsection \Assign\ + + text \ + The weakest precondition for assignments generates a lot of different cases. + However, usually only one of them is required for a given situation. + + The following rule eliminates the wrong cases by proving that they lead to a contradiction. + It requires two facts to be provided: + \<^item> @{term expr_rule}: should be a theorem which evaluates the expression part of the assignment + \<^item> @{term lexp_rule}: should be a theorem which evaluates the left hand side of the assignment + + + Both theorems should assume a corresponding loading of parameters and all declarations which happen before the assignment. + \ + + method vcg_insert_expr_lexp for ex::E and lv::L uses expr_rule lexp_rule = + match premises in + expr: "expr ex _ _ _ _ = _" and + lexp: "lexp lv _ _ _ _ = _" \ + \insert expr_rule[OF expr] lexp_rule[OF lexp]\ + + method vcg_insert_decl for ex::E and lv::L uses expr_rule lexp_rule = + match premises in + decl: "decl _ _ _ _ _ _ _ _ = _" (multi) \ + \vcg_insert_expr_lexp ex lv expr_rule:expr_rule[OF decl] lexp_rule:lexp_rule[OF decl]\ + \ _ \ + \vcg_insert_expr_lexp ex lv expr_rule:expr_rule lexp_rule:lexp_rule\ + + method vcg_insert_load for ex::E and lv::L uses expr_rule lexp_rule = + match premises in + load: "load _ _ _ _ _ _ _ _ _ _ _ = _" \ + \vcg_insert_decl ex lv expr_rule:expr_rule[OF load] lexp_rule:lexp_rule[OF load]\ + \ _ \ + \vcg_insert_decl ex lv expr_rule:expr_rule lexp_rule:lexp_rule\ + + method vcg_assign uses expr_rule lexp_rule = + match conclusion in + "wpS (stmt (ASSIGN lv ex) _ _) _ _ _" for lv ex \ + \rule wp_Assign; + (solve \(rule FalseE, simp, + (vcg_insert_load ex lv expr_rule:expr_rule lexp_rule:lexp_rule)), simp\ + | solve simp)?\, + simp + + subsection \Composition\ + + method vcg_comp = + rule wp_Comp; simp + + subsection \Blocks\ + + method vcg_block_some = + rule wp_blockSome; simp +end + +locale VCG = Calculus + + fixes pref::"Postcondition" + and postf::"Postcondition" + and pre::"Identifier \ Precondition" + and post::"Identifier \ Postcondition" +begin + +subsection \Transfer\ +text \ + The following rule can be used to verify an invariant for a transfer statement. + It requires four term parameters: + \<^item> @{term[show_types] "pref::Postcondition"}: Precondition for fallback method called internally + \<^item> @{term[show_types] "postf::Postcondition"}: Postcondition for fallback method called internally + \<^item> @{term[show_types] "pre::Identifier \ Precondition"}: Preconditions for internal methods + \<^item> @{term[show_types] "post::Identifier \ Postcondition"}: Postconditions for internal methods + + + In addition it requires 8 facts: + \<^item> @{term fallback_int}: verifies *postcondition* for body of fallback method invoked *internally*. + \<^item> @{term fallback_ext}: verifies *invariant* for body of fallback method invoked *externally*. + \<^item> @{term cases_ext}: performs case distinction over *external* methods of contract @{term ad}. + \<^item> @{term cases_int}: performs case distinction over *internal* methods of contract @{term ad}. + \<^item> @{term cases_fb}: performs case distinction over *fallback* method of contract @{term ad}. + \<^item> @{term different}: shows that address of environment is different from @{term ad}. + \<^item> @{term invariant}: shows that invariant holds *before* execution of transfer statement. + + + Finally it requires two lists of facts as parameters: + \<^item> @{term external}: verify that the invariant is preserved by the body of external methods. + \<^item> @{term internal}: verify that the postcondition holds after executing the body of internal methods. +\ + +method vcg_prep = + (rule allI)+, + rule impI, + (erule conjE)+ + +method vcg_body uses fallback_int fallback_ext cases_ext cases_int cases_fb = + (rule conjI)?, + match conclusion in + "Qe _ _ _" \ + \unfold Qe_def, + vcg_prep, + erule cases_ext; + (vcg_prep, + rule external; + solve \assumption | simp\)\ + \ "Qi _ _ _ _" \ + \unfold Qi_def, + vcg_prep, + erule cases_fb; + (vcg_prep, + rule internal; + solve \assumption | simp\)\ + \ "Qfi _ _ _ _" \ + \unfold Qfi_def, + rule allI, + rule impI, + rule cases_int; + (vcg_prep, + rule fallback_int; + solve \assumption | simp\)\ + \ "Qfe _ _ _" \ + \unfold Qfe_def, + rule allI, + rule impI, + rule cases_int; + (vcg_prep, + rule fallback_ext; + solve \assumption | simp\)\ + +method decl_load_rec for ad::Address and e::Environment uses eq decl load empty init = + match premises in + d: "decl _ _ _ _ _ _ _ (_, _, _, e') = Some (_, _, _, e)" for e'::Environment \ + \decl_load_rec ad e' eq:trans_sym[OF eq decl[OF d]] decl:decl load:load empty:empty init:init\ + \ l: "load _ _ _ (ffold (init members) (emptyEnv ad cname (address e') v) (fmdom members)) _ _ _ _ _ _ _ = Normal ((e, _, _, _), _)" for e'::Environment and v \ + \rule + trans[ + OF eq + trans[ + OF load[OF l] + trans[ + OF init[of (unchecked) members "emptyEnv ad cname (address e') v" "fmdom members"] + empty[of (unchecked) ad cname "address e'" v]]]]\ + +method sameaddr for ad::Address = + match conclusion in + "address e = ad" for e::Environment \ + \decl_load_rec ad e eq:refl[of "address e"] decl:decl_env[THEN conjunct1] load:msel_ssel_expr_load_rexp_gas(4)[THEN conjunct2, THEN conjunct1] init:ffold_init_ad empty:emptyEnv_address\ + +lemma eq_neq_eq_imp_neq: + "x = a \ b \ y \ a = b \ x \ y" by simp + +method sender for ad::Address = + match conclusion in + "adv \ ad" for adv::Address \ + \match premises in + a: "address e' \ ad" and e: "expr SENDER e _ _ _ = Normal ((KValue adv, Value TAddr), _)" for e::Environment and e'::Environment \ + \rule local.eq_neq_eq_imp_neq[OF expr_sender[OF e] a], + decl_load_rec ad e eq:refl[of "sender e"] decl:decl_env[THEN conjunct2, THEN conjunct1] load:msel_ssel_expr_load_rexp_gas(4)[THEN conjunct2, THEN conjunct2, THEN conjunct1] init:ffold_init_sender empty:emptyEnv_sender\\ + +method vcg_init for ad::Address uses invariant = + elims, + sameaddr ad, + sender ad, + (rule invariant; assumption) + +method vcg_transfer_type for ad::Address = + match premises in + type: "type (accounts _ ad) = Some (Contract cname)" and transfer: "Accounts.transfer _ ad _ _ = Some _" \ + \insert transfer_type_same[OF transfer, THEN trans, OF type], solve simp\ + +method vcg_transfer_ext for ad::Address + uses fallback_int fallback_ext cases_ext cases_int cases_fb invariant = + rule wp_transfer_ext[where pref = pref and postf = postf and pre = pre and post = post], + (vcg_body fallback_int:fallback_int fallback_ext:fallback_ext cases_ext:cases_ext cases_int:cases_int cases_fb:cases_fb)+, + vcg_transfer_type ad, + vcg_init ad invariant:invariant + +end + +end diff --git a/thys/Solidity/document/orcidlink.sty b/thys/Solidity/document/orcidlink.sty --- a/thys/Solidity/document/orcidlink.sty +++ b/thys/Solidity/document/orcidlink.sty @@ -1,63 +1,63 @@ -%% -%% This is file `orcidlink.sty', -%% generated with the docstrip utility. -%% -%% The original source files were: -%% -%% orcidlink.dtx (with options: `package') -%% -%% This is a generated file. -%% -%% Copyright (C) 2020 by Leo C. Stein -%% -------------------------------------------------------------------------- -%% This work may be distributed and/or modified under the -%% conditions of the LaTeX Project Public License, either version 1.3 -%% of this license or (at your option) any later version. -%% The latest version of this license is in -%% http://www.latex-project.org/lppl.txt -%% and version 1.3 or later is part of all distributions of LaTeX -%% version 2005/12/01 or later. -%% -\NeedsTeXFormat{LaTeX2e}[1994/06/01] -\ProvidesPackage{orcidlink} - [2021/06/11 v1.0.4 Linked ORCiD logo macro package] - -%% All I did was package up Milo's code on TeX.SE, -%% see https://tex.stackexchange.com/a/445583/34063 -\RequirePackage{hyperref} -\RequirePackage{tikz} - -\ProcessOptions\relax - -\usetikzlibrary{svg.path} - -\definecolor{orcidlogocol}{HTML}{A6CE39} -\tikzset{ - orcidlogo/.pic={ - \fill[orcidlogocol] svg{M256,128c0,70.7-57.3,128-128,128C57.3,256,0,198.7,0,128C0,57.3,57.3,0,128,0C198.7,0,256,57.3,256,128z}; - \fill[white] svg{M86.3,186.2H70.9V79.1h15.4v48.4V186.2z} - svg{M108.9,79.1h41.6c39.6,0,57,28.3,57,53.6c0,27.5-21.5,53.6-56.8,53.6h-41.8V79.1z M124.3,172.4h24.5c34.9,0,42.9-26.5,42.9-39.7c0-21.5-13.7-39.7-43.7-39.7h-23.7V172.4z} - svg{M88.7,56.8c0,5.5-4.5,10.1-10.1,10.1c-5.6,0-10.1-4.6-10.1-10.1c0-5.6,4.5-10.1,10.1-10.1C84.2,46.7,88.7,51.3,88.7,56.8z}; - } -} - -%% Reciprocal of the height of the svg whose source is above. The -%% original generates a 256pt high graphic; this macro holds 1/256. -\newcommand{\@OrigHeightRecip}{0.00390625} - -%% We will compute the current X height to make the logo the right height -\newlength{\@curXheight} - -\DeclareRobustCommand\orcidlink[1]{% -\texorpdfstring{% -\setlength{\@curXheight}{\fontcharht\font`X}% -\href{https://orcid.org/#1}{\XeTeXLinkBox{\mbox{% -\begin{tikzpicture}[yscale=-\@OrigHeightRecip*\@curXheight, -xscale=\@OrigHeightRecip*\@curXheight,transform shape] -\pic{orcidlogo}; -\end{tikzpicture}% -}}}}{}} - -\endinput -%% -%% End of file `orcidlink.sty'. +%% +%% This is file `orcidlink.sty', +%% generated with the docstrip utility. +%% +%% The original source files were: +%% +%% orcidlink.dtx (with options: `package') +%% +%% This is a generated file. +%% +%% Copyright (C) 2020 by Leo C. Stein +%% -------------------------------------------------------------------------- +%% This work may be distributed and/or modified under the +%% conditions of the LaTeX Project Public License, either version 1.3 +%% of this license or (at your option) any later version. +%% The latest version of this license is in +%% http://www.latex-project.org/lppl.txt +%% and version 1.3 or later is part of all distributions of LaTeX +%% version 2005/12/01 or later. +%% +\NeedsTeXFormat{LaTeX2e}[1994/06/01] +\ProvidesPackage{orcidlink} + [2021/06/11 v1.0.4 Linked ORCiD logo macro package] + +%% All I did was package up Milo's code on TeX.SE, +%% see https://tex.stackexchange.com/a/445583/34063 +\RequirePackage{hyperref} +\RequirePackage{tikz} + +\ProcessOptions\relax + +\usetikzlibrary{svg.path} + +\definecolor{orcidlogocol}{HTML}{A6CE39} +\tikzset{ + orcidlogo/.pic={ + \fill[orcidlogocol] svg{M256,128c0,70.7-57.3,128-128,128C57.3,256,0,198.7,0,128C0,57.3,57.3,0,128,0C198.7,0,256,57.3,256,128z}; + \fill[white] svg{M86.3,186.2H70.9V79.1h15.4v48.4V186.2z} + svg{M108.9,79.1h41.6c39.6,0,57,28.3,57,53.6c0,27.5-21.5,53.6-56.8,53.6h-41.8V79.1z M124.3,172.4h24.5c34.9,0,42.9-26.5,42.9-39.7c0-21.5-13.7-39.7-43.7-39.7h-23.7V172.4z} + svg{M88.7,56.8c0,5.5-4.5,10.1-10.1,10.1c-5.6,0-10.1-4.6-10.1-10.1c0-5.6,4.5-10.1,10.1-10.1C84.2,46.7,88.7,51.3,88.7,56.8z}; + } +} + +%% Reciprocal of the height of the svg whose source is above. The +%% original generates a 256pt high graphic; this macro holds 1/256. +\newcommand{\@OrigHeightRecip}{0.00390625} + +%% We will compute the current X height to make the logo the right height +\newlength{\@curXheight} + +\DeclareRobustCommand\orcidlink[1]{% +\texorpdfstring{% +\setlength{\@curXheight}{\fontcharht\font`X}% +\href{https://orcid.org/#1}{\XeTeXLinkBox{\mbox{% +\begin{tikzpicture}[yscale=-\@OrigHeightRecip*\@curXheight, +xscale=\@OrigHeightRecip*\@curXheight,transform shape] +\pic{orcidlogo}; +\end{tikzpicture}% +}}}}{}} + +\endinput +%% +%% End of file `orcidlink.sty'. diff --git a/thys/Solidity/document/root.bib b/thys/Solidity/document/root.bib --- a/thys/Solidity/document/root.bib +++ b/thys/Solidity/document/root.bib @@ -1,700 +1,700 @@ - -% Encoding: UTF-8 -@InProceedings{ Hodovan2018, - author = {Hodov\'{a}n, Ren\'{a}ta and Kiss, \'{A}kos and Gyim\'{o}thy, Tibor}, - booktitle = {Proceedings of the 9th ACM SIGSOFT International Workshop on Automating TEST Case Design, Selection, - and Evaluation}, - title = {{Grammarinator: A Grammar-Based Open Source Fuzzer}}, - year = 2018, - address = {New York, NY, USA}, - pages = {45{\^a}48}, - publisher = {Association for Computing Machinery}, - series = {A-TEST 2018}, - abstract = {Fuzzing, or random testing, is an increasingly popular testing technique. The power of the approach - lies in its ability to generate a large number of useful test cases without consuming expensive - manpower. Furthermore, because of the randomness, it can often produce unusual cases that would be - beyond the awareness of a human tester. In this paper, we present Grammarinator, a general purpose - test generator tool that is able to utilize existing parser grammars as models. Since the model can - act both as a parser and as a generator, the tool can provide the capabilities of both generation and - mutation-based fuzzers. The presented tool is actively used to test various JavaScript engines and has - found more than 100 unique issues.}, - doi = {10.1145/3278186.3278193}, - isbn = 9781450360531, - keywords = {grammars, security, random testing, fuzzing}, - location = {Lake Buena Vista, FL, USA}, - numpages = 4, - url = {https://doi.org/10.1145/3278186.3278193} -} - -@Misc{ Parr, - author = {Terence Parr}, - howpublished = {\url{https://www.antlr.org/index.html}}, - note = {{Accessed: 2021-05-01}}, - title = {ANTLR (ANother Tool for Language Recognition)} -} - -@Misc{ truffle, - author = {{ConsenSys Software Inc.}}, - howpublished = {\url{https://www.trufflesuite.com/truffle}}, - note = {{Accessed: 2021-05-01}}, - title = {Truffle} -} - -@Misc{ ganache, - author = {{ConsenSys Software Inc.}}, - howpublished = {\url{https://www.trufflesuite.com/docs/ganache/}}, - note = {{Accessed: 2021-05-01}}, - title = {Ganache} -} - -@Book{ Scott1971, - author = {Scott, Dana and Strachey, Christopher}, - publisher = {Oxford University Computing Laboratory, Programming Research Group Oxford}, - title = {Toward a mathematical semantics for computer languages}, - year = 1971, - volume = 1 -} - -@Book{ Scott1970, - author = {Scott, Dana}, - publisher = {Oxford University Computing Laboratory, Programming Research Group Oxford}, - title = {Outline of a mathematical theory of computation}, - year = 1970 -} - -@InProceedings{ Bertoni2013, - author = {Bertoni, Guido and Daemen, Joan and Peeters, Micha{\"e}l and Van Assche, Gilles}, - booktitle = {Advances in Cryptology -- EUROCRYPT 2013}, - title = {Keccak}, - year = 2013, - address = {Berlin, Heidelberg}, - editor = {Johansson, Thomas and Nguyen, Phong Q.}, - pages = {313--314}, - publisher = {Springer Berlin Heidelberg}, - abstract = {In October 2012, the American National Institute of Standards and Technology (NIST) announced the - selection of Keccak as the winner of the SHA-3 Cryptographic Hash Algorithm Competition [10,11]. This - concluded an open competition that was remarkable both for its magnitude and the involvement of the - cryptographic community. Public review is of paramount importance to increase the confidence in the - new standard and to favor its quick adoption. The SHA-3 competition explicitly took this into account - by giving open access to the candidate algorithms and everyone in the cryptographic community could - try to break them, compare their performance, or simply give comments.}, - isbn = {978-3-642-38348-9} -} - -@InProceedings{ Hildenbrandt2018, - author = {Hildenbrandt, Everett and Saxena, Manasvi and Rodrigues, Nishant and Zhu, Xiaoran and Daian, Philip - and Guth, Dwight and Moore, Brandon and Park, Daejun and Zhang, Yi and Stefanescu, Andrei and Rosu, - Grigore}, - booktitle = {2018 IEEE 31st Computer Security Foundations Symposium (CSF)}, - title = {KEVM: A Complete Formal Semantics of the Ethereum Virtual Machine}, - year = 2018, - pages = {204--217}, - doi = {10.1109/CSF.2018.00022} -} - -@Article{ Rosu2010, - author = {Grigore Ro{\`E}u and Traian Florin {\`E}erb{\"A}nut{\"A}}, - journal = {The Journal of Logic and Algebraic Programming}, - title = {An overview of the K semantic framework}, - year = 2010, - issn = {1567-8326}, - note = {Membrane computing and programming}, - number = 6, - pages = {397--434}, - volume = 79, - abstract = {K is an executable semantic framework in which programming languages, calculi, as well as type - systems or formal analysis tools can be defined, making use of configurations, computations and rules. - Configurations organize the system/program state in units called cells, which are labeled and can be - nested. Computations carry {\^a}computational meaning{\^a} as special nested list structures - sequentializing computational tasks, such as fragments of program; in particular, computations extend - the original language or calculus syntax. K (rewrite) rules generalize conventional rewrite rules by - making explicit which parts of the term they read, write, or do not care about. This distinction makes - K a suitable framework for defining truly concurrent languages or calculi, even in the presence of - sharing. Since computations can be handled like any other terms in a rewriting environment, that is, - they can be matched, moved from one place to another in the original term, modified, or even deleted, - K is particularly suitable for defining control-intensive language features such as abrupt - termination, exceptions, or call/cc. This paper gives an overview of the K framework: what it is, how - it can be used, and where it has been used so far. It also proposes and discusses the K definition of - Challenge, a programming language that aims to challenge and expose the limitations of existing - semantic frameworks.}, - doi = {https://doi.org/10.1016/j.jlap.2010.03.012}, - url = {https://www.sciencedirect.com/science/article/pii/S1567832610000160} -} - -@InProceedings{ Hirai2017, - author = {Hirai, Yoichi}, - booktitle = {Financial Cryptography and Data Security}, - title = {Defining the Ethereum Virtual Machine for Interactive Theorem Provers}, - year = 2017, - address = {Cham}, - editor = {Brenner, Michael and Rohloff, Kurt and Bonneau, Joseph and Miller, Andrew and Ryan, Peter Y.A. and - Teague, Vanessa and Bracciali, Andrea and Sala, Massimiliano and Pintore, Federico and Jakobsson, - Markus}, - pages = {520--535}, - publisher = {Springer International Publishing}, - abstract = {Smart contracts in Ethereum are executed by the Ethereum Virtual Machine (EVM). We defined EVM in - Lem, a language that can be compiled for a few interactive theorem provers. We tested our definition - against a standard test suite for Ethereum implementations. Using our definition, we proved some - safety properties of Ethereum smart contracts in an interactive theorem prover Isabelle/HOL. To our - knowledge, ours is the first formal EVM definition for smart contract verification that implements all - instructions. Our definition can serve as a basis for further analysis and generation of Ethereum - smart contracts.}, - isbn = {978-3-319-70278-0} -} - -@Book{ Nipkow2002, - author = {Tobias Nipkow and Lawrence C. Paulson and Markus Wenzel}, - publisher = {Springer}, - title = {Isabelle/HOL --- A Proof Assistant for Higher-Order Logic}, - year = 2002, - series = {LNCS}, - volume = 2283 -} - -@InProceedings{ perez.ea:smart:2021, - author = {Daniel Perez and Ben Livshits}, - title = {Smart Contract Vulnerabilities: Vulnerable Does Not Imply Exploited}, - booktitle = {30th {USENIX} Security Symposium ({USENIX} Security 21)}, - year = 2021, - url = {https://www.usenix.org/conference/usenixsecurity21/presentation/perez}, - publisher = {{USENIX} Association}, - month = aug -} - -@Article{ Yang2020, - author = {Yang, Zheng and Lei, Hang}, - journal = {Mathematical Problems in Engineering}, - title = {Lolisa: Formal Syntax and Semantics for a Subset of the Solidity Programming Language in Mathematical - Tool Coq}, - year = 2020, - issn = {1024-123X}, - pages = 6191537, - volume = 2020, - abstract = {The security of blockchain smart contracts is one of the most emerging issues of the greatest - interest for researchers. This article presents an intermediate specification language for the formal - verification of Ethereum-based smart contract in Coq, denoted as Lolisa. The formal syntax and - semantics of Lolisa contain a large subset of the Solidity programming language developed for the - Ethereum blockchain platform. To enhance type safety, the formal syntax of Lolisa adopts a stronger - static type system than Solidity. In addition, Lolisa includes a large subset of Solidity syntax - components as well as general-purpose programming language features. Therefore, Solidity programs can - be directly translated into Lolisa with line-by-line correspondence. Lolisa is inherently - generalizable and can be extended to express other programming languages. Finally, the syntax and - semantics of Lolisa have been encapsulated as an interpreter in mathematical tool Coq. Hence, smart - contracts written in Lolisa can be symbolically executed and verified in Coq.}, - publisher = {Hindawi}, - url = {https://doi.org/10.1155/2020/6191537} -} - -@Manual{ Coq, - title = {The Coq proof assistant reference manual}, - author = {The Coq development team}, - note = {Version 8.0}, - organization = {LogiCal Project}, - year = 2004, - url = {http://coq.inria.fr} -} - -@InProceedings{ Crosara2019, - author = {Marco Crosara and Gabriele Centurino and Vincenzo Arceri}, - booktitle = {Proceedings of The Eleventh International Conference on Advances in System Testing and Validation - Lifecycle (VALID)}, - title = {{Towards an Operational Semantics for Solidity}}, - year = 2019, - editor = {Jos van Rooyen and Samuele Buro and Marco Campion and Michele Pasqua}, - month = nov, - pages = {1--6}, - publisher = {IARIA} -} - -@InProceedings{ Jiao2020, - author = {Jiao, Jiao and Kan, Shuanglong and Lin, Shang-Wei and Sanan, David and Liu, Yang and Sun, Jun}, - booktitle = {2020 IEEE Symposium on Security and Privacy (SP)}, - title = {Semantic understanding of smart contracts: executable operational semantics of Solidity}, - year = 2020, - organization = {IEEE}, - pages = {1695--1712} -} - -@InProceedings{ Jiao2020a, - author = {Jiao, Jiao and Lin, Shang-Wei and Sun, Jun}, - booktitle = {Fundamental Approaches to Software Engineering}, - title = {A Generalized Formal Semantic Framework for Smart Contracts}, - year = 2020, - address = {Cham}, - editor = {Wehrheim, Heike and Cabot, Jordi}, - pages = {75--96}, - publisher = {Springer International Publishing}, - abstract = {Smart contracts can be regarded as one of the most popular blockchain-based applications. The - decentralized nature of the blockchain introduces vulnerabilities absent in other programs. - Furthermore, it is very difficult, if not impossible, to patch a smart contract after it has been - deployed. Therefore, smart contracts must be formally verified before they are deployed on the - blockchain to avoid attacks exploiting these vulnerabilities. There is a recent surge of interest in - analyzing and verifying smart contracts. While most of the existing works either focus on EVM bytecode - or translate Solidity contracts into programs in intermediate languages for analysis and verification, - we believe that a direct executable formal semantics of the high-level programming language of smart - contracts is necessary to guarantee the validity of the verification. In this work, we propose a - generalized formal semantic framework based on a general semantic model of smart contracts. - Furthermore, this framework can directly handle smart contracts written in different high-level - programming languages through semantic extensions and facilitates the formal verification of security - properties with the generated semantics.}, - isbn = {978-3-030-45234-6} -} - -@InProceedings{ Mavridou2019, - author = {Mavridou, Anastasia and Laszka, Aron and Stachtiari, Emmanouela and Dubey, Abhishek}, - booktitle = {Financial Cryptography and Data Security}, - title = {VeriSolid: Correct-by-Design Smart Contracts for Ethereum}, - year = 2019, - address = {Cham}, - editor = {Goldberg, Ian and Moore, Tyler}, - pages = {446--465}, - publisher = {Springer International Publishing}, - abstract = {The adoption of blockchain based distributed ledgers is growing fast due to their ability to provide - reliability, integrity, and auditability without trusted entities. One of the key capabilities of - these emerging platforms is the ability to create self-enforcing smart contracts. However, the - development of smart contracts has proven to be error-prone in practice, and as a result, contracts - deployed on public platforms are often riddled with security vulnerabilities. This issue is - exacerbated by the design of these platforms, which forbids updating contract code and rolling back - malicious transactions. In light of this, it is crucial to ensure that a smart contract is secure - before deploying it and trusting it with significant amounts of cryptocurrency. To this end, we - introduce the VeriSolid framework for the formal verification of contracts that are specified using a - transition-system based model with rigorous operational semantics. Our model-based approach allows - developers to reason about and verify contract behavior at a high level of abstraction. VeriSolid - allows the generation of Solidity code from the verified models, which enables the correct-by-design - development of smart{\^A}contracts.}, - isbn = {978-3-030-32101-7} -} - -@InProceedings{ Mavridou2018, - author = {Mavridou, Anastasia and Laszka, Aron}, - booktitle = {Principles of Security and Trust}, - title = {Tool Demonstration: FSolidM for Designing Secure Ethereum Smart Contracts}, - year = 2018, - address = {Cham}, - editor = {Bauer, Lujo and K{\"u}sters, Ralf}, - pages = {270--277}, - publisher = {Springer International Publishing}, - abstract = {Blockchain-based distributed computing platforms enable the trusted execution of - computation---defined in the form of smart contracts---without trusted agents. Smart contracts are - envisioned to have a variety of applications, ranging from financial to IoT asset tracking. - Unfortunately, the development of smart contracts has proven to be extremely error prone. In practice, - contracts are riddled with security vulnerabilities comprising a critical issue since bugs are by - design non-fixable and contracts may handle financial assets of significant value. To facilitate the - development of secure smart contracts, we have created the FSolidM framework, which allows developers - to define contracts as finite state machines (FSMs) with rigorous and clear semantics. FSolidM - provides an easy-to-use graphical editor for specifying FSMs, a code generator for creating Ethereum - smart contracts, and a set of plugins that developers may add to their FSMs to enhance security and - functionality.}, - isbn = {978-3-319-89722-6} -} - -@InProceedings{ Ahrendt2020, - author = {Ahrendt, Wolfgang and Bubel, Richard}, - booktitle = {Leveraging Applications of Formal Methods, Verification and Validation: Applications}, - title = {Functional Verification of Smart Contracts via Strong Data Integrity}, - year = 2020, - address = {Cham}, - editor = {Margaria, Tiziana and Steffen, Bernhard}, - pages = {9--24}, - publisher = {Springer International Publishing}, - abstract = {We present an invariant-based specification and verification methodology that allows us to - conveniently specify and verify strong data integrity properties for Solidity smart contracts. Our - approach is able to reason precisely about arbitrary usage of the contracts, which may include - re-entrance, a common security pitfall in smart contracts. We implemented the approach in a prototype - verification tool, called SolidiKeY, and applied it successfully to a number of smart contracts.}, - isbn = {978-3-030-61467-6} -} - -@Article{ Ahrendt2016, - author = {Ahrendt, Wolfgang and Beckert, Bernhard and Bubel, Richard and H{\"a}hnle, Reiner and Schmitt, Peter - H and Ulbrich, Mattias}, - journal = {Lecture Notes in Computer Science}, - title = {Deductive software verification--the key book}, - year = 2016, - volume = 10001, - publisher = {Springer} -} - -@InProceedings{ Crafa2020, - author = {Crafa, Silvia and Di Pirro, Matteo and Zucca, Elena}, - booktitle = {Financial Cryptography and Data Security}, - title = {Is Solidity Solid Enough?}, - year = 2020, - address = {Cham}, - editor = {Bracciali, Andrea and Clark, Jeremy and Pintore, Federico and R{\o}nne, Peter B. and Sala, - Massimiliano}, - pages = {138--153}, - publisher = {Springer International Publishing}, - abstract = {We introduce Featherweight Solidity, a calculus formalizing the core features of the Solidity - language, thus providing a fundamental step to reason about safety properties of smart contracts' - source code. The formalization includes a static type system that represents the foundation of the - Solidity compiler. We show that it prevents some errors whereas many others, such as accesses to a non - existing function or state variable, are only detected at runtime and cause interruption and - rolling-back of transactions. We then propose a refinement of the type system that is retro-compatible - with original Solidity code, and statically captures more errors, such as unsafe casts and unsafe - call-back expressions.}, - isbn = {978-3-030-43725-1} -} - -@Misc{ Remix, - author = {Online}, - howpublished = {\url{https://remix-ide.readthedocs.io/en/latest}}, - title = {Remix - Solidity IDE} -} - -@InProceedings{ Bartoletti2019, - author = {Bartoletti, Massimo and Galletta, Letterio and Murgia, Maurizio}, - booktitle = {Data Privacy Management, Cryptocurrencies and Blockchain Technology}, - title = {A Minimal Core Calculus for Solidity Contracts}, - year = 2019, - address = {Cham}, - editor = {P{\'e}rez-Sol{\`a}, Cristina and Navarro-Arribas, Guillermo and Biryukov, Alex and Garcia-Alfaro, - Joaquin}, - pages = {233--243}, - publisher = {Springer International Publishing}, - abstract = {The Ethereum platform supports the decentralized execution of smart contracts, i.e. computer programs - that transfer digital assets between users. The most common language used to develop these contracts - is Solidity, a Javascript-like language which compiles into EVM bytecode, the language actually - executed by Ethereum nodes. While much research has addressed the formalisation of the semantics of - EVM bytecode, relatively little attention has been devoted to that of Solidity. In this paper we - propose a minimal calculus for Solidity contracts, which extends an imperative core with a single - primitive to transfer currency and invoke contract procedures. We build upon this formalisation to - give semantics to the Ethereum blockchain. We show our calculus expressive enough to reason about some - typical quirks of Solidity, like e.g. re-entrancy.}, - isbn = {978-3-030-31500-9} -} - -@InProceedings{ Zakrzewski2018, - author = {Jakub Zakrzewski}, - booktitle = {Verified Software. Theories, Tools, and Experiments - 10th International Conference, {VSTTE} 2018, - Oxford, UK, July 18-19, 2018, Revised Selected Papers}, - title = {Towards Verification of Ethereum Smart Contracts: {A} Formalization of Core of Solidity}, - year = 2018, - editor = {Ruzica Piskac and Philipp R{\"{u}}mmer}, - pages = {229--247}, - publisher = {Springer}, - series = {Lecture Notes in Computer Science}, - volume = 11294, - bibsource = {dblp computer science bibliography, https://dblp.org}, - biburl = {https://dblp.org/rec/conf/vstte/Zakrzewski18.bib}, - doi = {10.1007/978-3-030-03592-1\_13}, - timestamp = {Tue, 14 May 2019 10:00:49 +0200}, - url = {https://doi.org/10.1007/978-3-030-03592-1\_13} -} - -@InProceedings{ Hajdu2020, - author = {Hajdu, {\'A}kos and Jovanovi{\'{c}}, Dejan}, - booktitle = {Verified Software. Theories, Tools, and Experiments}, - title = {solc-verify: A Modular Verifier for Solidity Smart Contracts}, - year = 2020, - address = {Cham}, - editor = {Chakraborty, Supratik and Navas, Jorge A.}, - pages = {161--179}, - publisher = {Springer International Publishing}, - abstract = {We present solc-verify, a source-level verification tool for Ethereum smart contracts. solc-verify - takes smart contracts written in Solidity and discharges verification conditions using modular program - analysis and SMT solvers. Built on top of the Solidity compiler, solc-verify reasons at the level of - the contract source code, as opposed to the more common approaches that operate at the level of - Ethereum bytecode. This enables solc-verify to effectively reason about high-level contract properties - while modeling low-level language semantics precisely. The properties, such as contract invariants, - loop invariants, and function pre- and post-conditions, can be provided as annotations in the code by - the developer. This enables automated, yet user-friendly formal verification for smart contracts. We - demonstrate solc-verify by examining real-world examples where our tool can effectively find bugs and - prove correctness of non-trivial properties with minimal user effort.}, - isbn = {978-3-030-41600-3} -} - -@InProceedings{ Amani2018, - author = {Amani, Sidney and B\'{e}gel, Myriam and Bortin, Maksym and Staples, Mark}, - booktitle = {Proceedings of the 7th ACM SIGPLAN International Conference on Certified Programs and Proofs}, - title = {Towards Verifying Ethereum Smart Contract Bytecode in Isabelle/HOL}, - year = 2018, - address = {New York, NY, USA}, - pages = {66{\^a}77}, - publisher = {Association for Computing Machinery}, - series = {CPP 2018}, - abstract = {Blockchain technology has increasing attention in research and across many industries. The Ethereum - blockchain offers smart contracts, which are small programs defined, executed, and recorded as - transactions in the blockchain transaction history. These smart contracts run on the Ethereum Virtual - Machine (EVM) and can be used to encode agreements, transfer assets, and enforce integrity conditions - in relationships between parties. Smart contracts can carry financial value, and are increasingly used - for safety-, security-, or mission-critical purposes. Errors in smart contracts have led and will lead - to loss or harm. Formal verification can provide the highest level of confidence about the correct - behaviour of smart contracts. In this paper we extend an existing EVM formalisation in Isabelle/HOL by - a sound program logic at the level of bytecode. We structure bytecode sequences into blocks of - straight-line code and create a program logic to reason about these. This abstraction is a step - towards control of the cost and complexity of formal verification of EVM smart contracts.}, - doi = {10.1145/3167084}, - isbn = 9781450355865, - keywords = {blockchain, formal verification, smart contracts, Ethereum, Isabelle/HOL}, - location = {Los Angeles, CA, USA}, - numpages = 12, - url = {https://doi.org/10.1145/3167084} -} - -@Misc{ Wood2014, - author = {Wood, Gavin and others}, - journal = {Ethereum project yellow paper}, - title = {Ethereum: A secure decentralised generalised transaction ledger}, - year = 2022, - note = {Berlin Version 3078285 -- 2022-07-13. \url{https://ethereum.github.io/yellowpaper/paper.pdf}}, - pages = {1--32}, -} - -@Misc{ Solidity, - author = {Online}, - howpublished = {\url{https://solidity.readthedocs.io/en/latest}}, - title = {Solidity documentation} -} - -@InProceedings{ Grishchenko2018, - author = {Grishchenko, Ilya and Maffei, Matteo and Schneidewind, Clara}, - booktitle = {Principles of Security and Trust}, - title = {A Semantic Framework for the Security Analysis of Ethereum Smart Contracts}, - year = 2018, - address = {Cham}, - editor = {Bauer, Lujo and K{\"u}sters, Ralf}, - pages = {243--269}, - publisher = {Springer International Publishing}, - abstract = {Smart contracts are programs running on cryptocurrency (e.g., Ethereum) blockchains, whose popularity - stem from the possibility to perform financial transactions, such as payments and auctions, in a - distributed environment without need for any trusted third party. Given their financial nature, bugs - or vulnerabilities in these programs may lead to catastrophic consequences, as witnessed by recent - attacks. Unfortunately, programming smart contracts is a delicate task that requires strong expertise: - Ethereum smart contracts are written in Solidity, a dedicated language resembling JavaScript, and - shipped over the blockchain in the EVM bytecode format. In order to rigorously verify the security of - smart contracts, it is of paramount importance to formalize their semantics as well as the security - properties of interest, in particular at the level of the bytecode being executed.}, - isbn = {978-3-319-89722-6} -} - -@Article{ Wood2014a, - author = {Wood, Gavin and others}, - journal = {Ethereum project yellow paper}, - title = {Ethereum: A secure decentralised generalised transaction ledger}, - year = 2014, - number = 2014, - pages = {1--32}, - volume = 151 -} - -@InProceedings{ Swamy2016, - author = {Swamy, Nikhil and Hri\c{t}cu, C\u{a}t\u{a}lin and Keller, Chantal and Rastogi, Aseem and - Delignat-Lavaud, Antoine and Forest, Simon and Bhargavan, Karthikeyan and Fournet, C\'{e}dric and - Strub, Pierre-Yves and Kohlweiss, Markulf and Zinzindohoue, Jean-Karim and Zanella-B\'{e}guelin, Santiago}, - booktitle = {Proceedings of the 43rd Annual ACM SIGPLAN-SIGACT Symposium on Principles of Programming Languages}, - title = {Dependent Types and Multi-Monadic Effects in F*}, - year = 2016, - address = {New York, NY, USA}, - pages = {256{\^a}270}, - publisher = {Association for Computing Machinery}, - series = {POPL '16}, - abstract = {We present a new, completely redesigned, version of F*, a language that works both as a proof - assistant as well as a general-purpose, verification-oriented, effectful programming language. In - support of these complementary roles, F* is a dependently typed, higher-order, call-by-value language - with _primitive_ effects including state, exceptions, divergence and IO. Although primitive, - programmers choose the granularity at which to specify effects by equipping each effect with a - monadic, predicate transformer semantics. F* uses this to efficiently compute weakest preconditions - and discharges the resulting proof obligations using a combination of SMT solving and manual proofs. - Isolated from the effects, the core of F* is a language of pure functions used to write specifications - and proof terms---its consistency is maintained by a semantic termination check based on a - well-founded order. We evaluate our design on more than 55,000 lines of F* we have authored in the - last year, focusing on three main case studies. Showcasing its use as a general-purpose programming - language, F* is programmed (but not verified) in F*, and bootstraps in both OCaml and F#. Our - experience confirms F*'s pay-as-you-go cost model: writing idiomatic ML-like code with no finer - specifications imposes no user burden. As a verification-oriented language, our most significant - evaluation of F* is in verifying several key modules in an implementation of the TLS-1.2 protocol - standard. For the modules we considered, we are able to prove more properties, with fewer annotations - using F* than in a prior verified implementation of TLS-1.2. Finally, as a proof assistant, we discuss - our use of F* in mechanizing the metatheory of a range of lambda calculi, starting from the simply - typed lambda calculus to System F-omega and even micro-F*, a sizeable fragment of F* itself---these - proofs make essential use of F*'s flexible combination of SMT automation and constructive proofs, - enabling a tactic-free style of programming and proving at a relatively large scale.}, - doi = {10.1145/2837614.2837655}, - isbn = 9781450335492, - keywords = {effectful programming, verification, proof assistants}, - location = {St. Petersburg, FL, USA}, - numpages = 15, - url = {https://doi.org/10.1145/2837614.2837655} -} - -@Article{ 10.1145/2914770.2837655, - author = {Swamy, Nikhil and Hri\c{t}cu, C\u{a}t\u{a}lin and Keller, Chantal and Rastogi, Aseem and - Delignat-Lavaud, Antoine and Forest, Simon and Bhargavan, Karthikeyan and Fournet, C\'{e}dric and - Strub, Pierre-Yves and Kohlweiss, Markulf and Zinzindohoue, Jean-Karim and Zanella-B\'{e}guelin, Santiago}, - journal = {SIGPLAN Not.}, - title = {Dependent Types and Multi-Monadic Effects in F*}, - year = 2016, - issn = {0362-1340}, - month = jan, - number = 1, - pages = {256{\^a}270}, - volume = 51, - abstract = {We present a new, completely redesigned, version of F*, a language that works both as a proof - assistant as well as a general-purpose, verification-oriented, effectful programming language. In - support of these complementary roles, F* is a dependently typed, higher-order, call-by-value language - with _primitive_ effects including state, exceptions, divergence and IO. Although primitive, - programmers choose the granularity at which to specify effects by equipping each effect with a - monadic, predicate transformer semantics. F* uses this to efficiently compute weakest preconditions - and discharges the resulting proof obligations using a combination of SMT solving and manual proofs. - Isolated from the effects, the core of F* is a language of pure functions used to write specifications - and proof terms---its consistency is maintained by a semantic termination check based on a - well-founded order. We evaluate our design on more than 55,000 lines of F* we have authored in the - last year, focusing on three main case studies. Showcasing its use as a general-purpose programming - language, F* is programmed (but not verified) in F*, and bootstraps in both OCaml and F#. Our - experience confirms F*'s pay-as-you-go cost model: writing idiomatic ML-like code with no finer - specifications imposes no user burden. As a verification-oriented language, our most significant - evaluation of F* is in verifying several key modules in an implementation of the TLS-1.2 protocol - standard. For the modules we considered, we are able to prove more properties, with fewer annotations - using F* than in a prior verified implementation of TLS-1.2. Finally, as a proof assistant, we discuss - our use of F* in mechanizing the metatheory of a range of lambda calculi, starting from the simply - typed lambda calculus to System F-omega and even micro-F*, a sizeable fragment of F* itself---these - proofs make essential use of F*'s flexible combination of SMT automation and constructive proofs, - enabling a tactic-free style of programming and proving at a relatively large scale.}, - address = {New York, NY, USA}, - doi = {10.1145/2914770.2837655}, - issue_date = {January 2016}, - keywords = {proof assistants, verification, effectful programming}, - numpages = 15, - publisher = {Association for Computing Machinery}, - url = {https://doi.org/10.1145/2914770.2837655} -} - -@Article{ Mulligan2014, - author = {Mulligan, Dominic P. and Owens, Scott and Gray, Kathryn E. and Ridge, Tom and Sewell, Peter}, - journal = {SIGPLAN Not.}, - title = {Lem: Reusable Engineering of Real-World Semantics}, - year = 2014, - issn = {0362-1340}, - month = aug, - number = 9, - pages = {175{\^a}188}, - volume = 49, - abstract = {Recent years have seen remarkable successes in rigorous engineering: using mathematically rigorous - semantic models (not just idealised calculi) of real-world processors, programming languages, - protocols, and security mechanisms, for testing, proof, analysis, and design. Building these models is - challenging, requiring experimentation, dialogue with vendors or standards bodies, and validation; - their scale adds engineering issues akin to those of programming to the task of writing clear and - usable mathematics. But language and tool support for specification is lacking. Proof assistants can - be used but bring their own difficulties, and a model produced in one, perhaps requiring many - person-years effort and maintained over an extended period, cannot be used by those familiar with - another.We introduce Lem, a language for engineering reusable large-scale semantic models. The Lem - design takes inspiration both from functional programming languages and from proof assistants, and Lem - definitions are translatable into OCaml for testing, Coq, HOL4, and Isabelle/HOL for proof, and LaTeX - and HTML for presentation. This requires a delicate balance of expressiveness, careful library design, - and implementation of transformations - akin to compilation, but subject to the constraint of - producing usable and human-readable code for each target. Lem's effectiveness is demonstrated by its - use in practice.}, - address = {New York, NY, USA}, - doi = {10.1145/2692915.2628143}, - issue_date = {September 2014}, - keywords = {proof assistants, real-world semantics, lem, specification languages}, - numpages = 14, - publisher = {Association for Computing Machinery}, - url = {https://doi.org/10.1145/2692915.2628143} -} - -@InProceedings{ 10.1145/2628136.2628143, - author = {Mulligan, Dominic P. and Owens, Scott and Gray, Kathryn E. and Ridge, Tom and Sewell, Peter}, - booktitle = {Proceedings of the 19th ACM SIGPLAN International Conference on Functional Programming}, - title = {Lem: Reusable Engineering of Real-World Semantics}, - year = 2014, - address = {New York, NY, USA}, - pages = {175{\^a}188}, - publisher = {Association for Computing Machinery}, - series = {ICFP '14}, - abstract = {Recent years have seen remarkable successes in rigorous engineering: using mathematically rigorous - semantic models (not just idealised calculi) of real-world processors, programming languages, - protocols, and security mechanisms, for testing, proof, analysis, and design. Building these models is - challenging, requiring experimentation, dialogue with vendors or standards bodies, and validation; - their scale adds engineering issues akin to those of programming to the task of writing clear and - usable mathematics. But language and tool support for specification is lacking. Proof assistants can - be used but bring their own difficulties, and a model produced in one, perhaps requiring many - person-years effort and maintained over an extended period, cannot be used by those familiar with - another.We introduce Lem, a language for engineering reusable large-scale semantic models. The Lem - design takes inspiration both from functional programming languages and from proof assistants, and Lem - definitions are translatable into OCaml for testing, Coq, HOL4, and Isabelle/HOL for proof, and LaTeX - and HTML for presentation. This requires a delicate balance of expressiveness, careful library design, - and implementation of transformations - akin to compilation, but subject to the constraint of - producing usable and human-readable code for each target. Lem's effectiveness is demonstrated by its - use in practice.}, - doi = {10.1145/2628136.2628143}, - isbn = 9781450328739, - keywords = {real-world semantics, lem, specification languages, proof assistants}, - location = {Gothenburg, Sweden}, - numpages = 14, - url = {https://doi.org/10.1145/2628136.2628143} -} - -@Manual{ solidity-0.8.4, - title = {Solidity (Version 0.8.4)}, - howpublished = {\url{https://docs.soliditylang.org/en/v0.8.4/}} -} - -@TechReport{ wood:ethereum:2021, - author = {Gavin Wood}, - title = {Ethereum: A Secure Decentralised Generalised Transation Ledger (Version 2021-04-21)}, - optnote = {\url{https://ethereum.github.io/yellowpaper/paper.pdf}} -} - -@Misc{ coinmarket, - title = {The {Bitcon} Market Capitalisation.}, - url = {https://coinmarketcap.com/currencies/bitcoin/}, - note = {Last checked on 2021-05-04.} -} - -@InCollection{ marmsoler.ea:conformance:2022, - abstract = {A common problem in verification is to ensure that the formal specification models the real-world - system, i.e., the implementation, faithfully. Testing is a technique that can help to bridge the gap - between a formal specification and its implementation.\\\\Fuzzing in general and grammar-based fuzzing - in particular are successfully used for finding bugs in implementations. Traditional fuzzing - applications rely on an implicit test specification that informally can be described as ``the program - under test does not crash''.\\\\In this paper, we present an approach using grammar-based fuzzing to - ensure the conformance of a formal specification, namely the formal semantics of the Solidity - Programming language, to a real-world implementation. For this, we derive an executable test-oracle - from the formal semantics of Solidity in Isabelle/HOL. The derived test oracle is used during the - fuzzing of the implementation to validate that the formal semantics and the implementation are in - conformance.}, - address = {Heidelberg}, - author = {Diego Marmsoler and Achim D. Brucker}, - booktitle = {{TAP} 2022: Tests And Proofs}, - editor = {Laura Kovacs and Karl Meinke}, - isbn = {978-3-642-38915-3}, - keywords = {Conformance Testing, Fuzzing, Verification, Solidity}, - language = {USenglish}, - location = {Nantes, France}, - pdf = {https://www.brucker.ch/bibliography/download/2022/marmsoler.ea-conformance-2022.pdf}, - publisher = {Springer-Verlag}, - series = {Lecture Notes in Computer Science}, - title = {Conformance Testing of Formal Semantics using Grammar-based Fuzzing}, - url = {https://www.brucker.ch/bibliography/abstract/marmsoler.ea-conformance-2022}, - year = 2022 -} - -@InCollection{ marmsoler.ea:solidity-semantics:2021, - abstract = {Smart contracts are programs, usually automating legal agreements such as financial transactions. - Thus, bugs in smart contracts can lead to large financial losses. For example, an incorrectly - initialized contract was the root cause of the Parity Wallet bug that made USD 280mil worth of Ether - inaccessible. Ether is the cryptocurrency of the Ethereum blockchain that uses Solidity for expressing - smart contracts.\\\\In this paper, we address this problem by presenting an executable denotational - semantics for Solidity in the interactive theorem prover Isabelle/HOL. This formal semantics builds - the foundation of an interactive program verification environment for Solidity programs and allows for - inspecting Solidity programs by (symbolic) execution. We combine the latter with grammar-based fuzzing - to ensure that our formal semantics complies to the Solidity implementation on the Ethereum - Blockchain. Finally, we demonstrate the formal verification of Solidity programs by two examples: - constant folding and memory optimization.}, - address = {Heidelberg}, - author = {Diego Marmsoler and Achim D. Brucker}, - booktitle = {Software Engineering and Formal Methods (SEFM)}, - editor = {Radu Calinescu and Corina Pasareanu}, - isbn = {3-540-25109-X}, - keywords = {Solidity, Denotational Semantics, Isabelle/HOL, Gas Optimization}, - language = {USenglish}, - pdf = {https://www.brucker.ch/bibliography/download/2021/marmsoler.ea-solidity-semantics-2021.pdf}, - publisher = {Springer-Verlag}, - series = {Lecture Notes in Computer Science}, - title = {A Denotational Semantics of {Solidity} in {Isabelle/HOL}}, - url = {https://www.brucker.ch/bibliography/abstract/marmsoler.ea-solidity-semantics-2021}, - year = 2021 -} -@COMMENT{jabref-meta: databaseType:bibtex;} + +% Encoding: UTF-8 +@InProceedings{ Hodovan2018, + author = {Hodov\'{a}n, Ren\'{a}ta and Kiss, \'{A}kos and Gyim\'{o}thy, Tibor}, + booktitle = {Proceedings of the 9th ACM SIGSOFT International Workshop on Automating TEST Case Design, Selection, + and Evaluation}, + title = {{Grammarinator: A Grammar-Based Open Source Fuzzer}}, + year = 2018, + address = {New York, NY, USA}, + pages = {45{\^a}48}, + publisher = {Association for Computing Machinery}, + series = {A-TEST 2018}, + abstract = {Fuzzing, or random testing, is an increasingly popular testing technique. The power of the approach + lies in its ability to generate a large number of useful test cases without consuming expensive + manpower. Furthermore, because of the randomness, it can often produce unusual cases that would be + beyond the awareness of a human tester. In this paper, we present Grammarinator, a general purpose + test generator tool that is able to utilize existing parser grammars as models. Since the model can + act both as a parser and as a generator, the tool can provide the capabilities of both generation and + mutation-based fuzzers. The presented tool is actively used to test various JavaScript engines and has + found more than 100 unique issues.}, + doi = {10.1145/3278186.3278193}, + isbn = 9781450360531, + keywords = {grammars, security, random testing, fuzzing}, + location = {Lake Buena Vista, FL, USA}, + numpages = 4, + url = {https://doi.org/10.1145/3278186.3278193} +} + +@Misc{ Parr, + author = {Terence Parr}, + howpublished = {\url{https://www.antlr.org/index.html}}, + note = {{Accessed: 2021-05-01}}, + title = {ANTLR (ANother Tool for Language Recognition)} +} + +@Misc{ truffle, + author = {{ConsenSys Software Inc.}}, + howpublished = {\url{https://www.trufflesuite.com/truffle}}, + note = {{Accessed: 2021-05-01}}, + title = {Truffle} +} + +@Misc{ ganache, + author = {{ConsenSys Software Inc.}}, + howpublished = {\url{https://www.trufflesuite.com/docs/ganache/}}, + note = {{Accessed: 2021-05-01}}, + title = {Ganache} +} + +@Book{ Scott1971, + author = {Scott, Dana and Strachey, Christopher}, + publisher = {Oxford University Computing Laboratory, Programming Research Group Oxford}, + title = {Toward a mathematical semantics for computer languages}, + year = 1971, + volume = 1 +} + +@Book{ Scott1970, + author = {Scott, Dana}, + publisher = {Oxford University Computing Laboratory, Programming Research Group Oxford}, + title = {Outline of a mathematical theory of computation}, + year = 1970 +} + +@InProceedings{ Bertoni2013, + author = {Bertoni, Guido and Daemen, Joan and Peeters, Micha{\"e}l and Van Assche, Gilles}, + booktitle = {Advances in Cryptology -- EUROCRYPT 2013}, + title = {Keccak}, + year = 2013, + address = {Berlin, Heidelberg}, + editor = {Johansson, Thomas and Nguyen, Phong Q.}, + pages = {313--314}, + publisher = {Springer Berlin Heidelberg}, + abstract = {In October 2012, the American National Institute of Standards and Technology (NIST) announced the + selection of Keccak as the winner of the SHA-3 Cryptographic Hash Algorithm Competition [10,11]. This + concluded an open competition that was remarkable both for its magnitude and the involvement of the + cryptographic community. Public review is of paramount importance to increase the confidence in the + new standard and to favor its quick adoption. The SHA-3 competition explicitly took this into account + by giving open access to the candidate algorithms and everyone in the cryptographic community could + try to break them, compare their performance, or simply give comments.}, + isbn = {978-3-642-38348-9} +} + +@InProceedings{ Hildenbrandt2018, + author = {Hildenbrandt, Everett and Saxena, Manasvi and Rodrigues, Nishant and Zhu, Xiaoran and Daian, Philip + and Guth, Dwight and Moore, Brandon and Park, Daejun and Zhang, Yi and Stefanescu, Andrei and Rosu, + Grigore}, + booktitle = {2018 IEEE 31st Computer Security Foundations Symposium (CSF)}, + title = {KEVM: A Complete Formal Semantics of the Ethereum Virtual Machine}, + year = 2018, + pages = {204--217}, + doi = {10.1109/CSF.2018.00022} +} + +@Article{ Rosu2010, + author = {Grigore Ro{\`E}u and Traian Florin {\`E}erb{\"A}nut{\"A}}, + journal = {The Journal of Logic and Algebraic Programming}, + title = {An overview of the K semantic framework}, + year = 2010, + issn = {1567-8326}, + note = {Membrane computing and programming}, + number = 6, + pages = {397--434}, + volume = 79, + abstract = {K is an executable semantic framework in which programming languages, calculi, as well as type + systems or formal analysis tools can be defined, making use of configurations, computations and rules. + Configurations organize the system/program state in units called cells, which are labeled and can be + nested. Computations carry {\^a}computational meaning{\^a} as special nested list structures + sequentializing computational tasks, such as fragments of program; in particular, computations extend + the original language or calculus syntax. K (rewrite) rules generalize conventional rewrite rules by + making explicit which parts of the term they read, write, or do not care about. This distinction makes + K a suitable framework for defining truly concurrent languages or calculi, even in the presence of + sharing. Since computations can be handled like any other terms in a rewriting environment, that is, + they can be matched, moved from one place to another in the original term, modified, or even deleted, + K is particularly suitable for defining control-intensive language features such as abrupt + termination, exceptions, or call/cc. This paper gives an overview of the K framework: what it is, how + it can be used, and where it has been used so far. It also proposes and discusses the K definition of + Challenge, a programming language that aims to challenge and expose the limitations of existing + semantic frameworks.}, + doi = {https://doi.org/10.1016/j.jlap.2010.03.012}, + url = {https://www.sciencedirect.com/science/article/pii/S1567832610000160} +} + +@InProceedings{ Hirai2017, + author = {Hirai, Yoichi}, + booktitle = {Financial Cryptography and Data Security}, + title = {Defining the Ethereum Virtual Machine for Interactive Theorem Provers}, + year = 2017, + address = {Cham}, + editor = {Brenner, Michael and Rohloff, Kurt and Bonneau, Joseph and Miller, Andrew and Ryan, Peter Y.A. and + Teague, Vanessa and Bracciali, Andrea and Sala, Massimiliano and Pintore, Federico and Jakobsson, + Markus}, + pages = {520--535}, + publisher = {Springer International Publishing}, + abstract = {Smart contracts in Ethereum are executed by the Ethereum Virtual Machine (EVM). We defined EVM in + Lem, a language that can be compiled for a few interactive theorem provers. We tested our definition + against a standard test suite for Ethereum implementations. Using our definition, we proved some + safety properties of Ethereum smart contracts in an interactive theorem prover Isabelle/HOL. To our + knowledge, ours is the first formal EVM definition for smart contract verification that implements all + instructions. Our definition can serve as a basis for further analysis and generation of Ethereum + smart contracts.}, + isbn = {978-3-319-70278-0} +} + +@Book{ Nipkow2002, + author = {Tobias Nipkow and Lawrence C. Paulson and Markus Wenzel}, + publisher = {Springer}, + title = {Isabelle/HOL --- A Proof Assistant for Higher-Order Logic}, + year = 2002, + series = {LNCS}, + volume = 2283 +} + +@InProceedings{ perez.ea:smart:2021, + author = {Daniel Perez and Ben Livshits}, + title = {Smart Contract Vulnerabilities: Vulnerable Does Not Imply Exploited}, + booktitle = {30th {USENIX} Security Symposium ({USENIX} Security 21)}, + year = 2021, + url = {https://www.usenix.org/conference/usenixsecurity21/presentation/perez}, + publisher = {{USENIX} Association}, + month = aug +} + +@Article{ Yang2020, + author = {Yang, Zheng and Lei, Hang}, + journal = {Mathematical Problems in Engineering}, + title = {Lolisa: Formal Syntax and Semantics for a Subset of the Solidity Programming Language in Mathematical + Tool Coq}, + year = 2020, + issn = {1024-123X}, + pages = 6191537, + volume = 2020, + abstract = {The security of blockchain smart contracts is one of the most emerging issues of the greatest + interest for researchers. This article presents an intermediate specification language for the formal + verification of Ethereum-based smart contract in Coq, denoted as Lolisa. The formal syntax and + semantics of Lolisa contain a large subset of the Solidity programming language developed for the + Ethereum blockchain platform. To enhance type safety, the formal syntax of Lolisa adopts a stronger + static type system than Solidity. In addition, Lolisa includes a large subset of Solidity syntax + components as well as general-purpose programming language features. Therefore, Solidity programs can + be directly translated into Lolisa with line-by-line correspondence. Lolisa is inherently + generalizable and can be extended to express other programming languages. Finally, the syntax and + semantics of Lolisa have been encapsulated as an interpreter in mathematical tool Coq. Hence, smart + contracts written in Lolisa can be symbolically executed and verified in Coq.}, + publisher = {Hindawi}, + url = {https://doi.org/10.1155/2020/6191537} +} + +@Manual{ Coq, + title = {The Coq proof assistant reference manual}, + author = {The Coq development team}, + note = {Version 8.0}, + organization = {LogiCal Project}, + year = 2004, + url = {http://coq.inria.fr} +} + +@InProceedings{ Crosara2019, + author = {Marco Crosara and Gabriele Centurino and Vincenzo Arceri}, + booktitle = {Proceedings of The Eleventh International Conference on Advances in System Testing and Validation + Lifecycle (VALID)}, + title = {{Towards an Operational Semantics for Solidity}}, + year = 2019, + editor = {Jos van Rooyen and Samuele Buro and Marco Campion and Michele Pasqua}, + month = nov, + pages = {1--6}, + publisher = {IARIA} +} + +@InProceedings{ Jiao2020, + author = {Jiao, Jiao and Kan, Shuanglong and Lin, Shang-Wei and Sanan, David and Liu, Yang and Sun, Jun}, + booktitle = {2020 IEEE Symposium on Security and Privacy (SP)}, + title = {Semantic understanding of smart contracts: executable operational semantics of Solidity}, + year = 2020, + organization = {IEEE}, + pages = {1695--1712} +} + +@InProceedings{ Jiao2020a, + author = {Jiao, Jiao and Lin, Shang-Wei and Sun, Jun}, + booktitle = {Fundamental Approaches to Software Engineering}, + title = {A Generalized Formal Semantic Framework for Smart Contracts}, + year = 2020, + address = {Cham}, + editor = {Wehrheim, Heike and Cabot, Jordi}, + pages = {75--96}, + publisher = {Springer International Publishing}, + abstract = {Smart contracts can be regarded as one of the most popular blockchain-based applications. The + decentralized nature of the blockchain introduces vulnerabilities absent in other programs. + Furthermore, it is very difficult, if not impossible, to patch a smart contract after it has been + deployed. Therefore, smart contracts must be formally verified before they are deployed on the + blockchain to avoid attacks exploiting these vulnerabilities. There is a recent surge of interest in + analyzing and verifying smart contracts. While most of the existing works either focus on EVM bytecode + or translate Solidity contracts into programs in intermediate languages for analysis and verification, + we believe that a direct executable formal semantics of the high-level programming language of smart + contracts is necessary to guarantee the validity of the verification. In this work, we propose a + generalized formal semantic framework based on a general semantic model of smart contracts. + Furthermore, this framework can directly handle smart contracts written in different high-level + programming languages through semantic extensions and facilitates the formal verification of security + properties with the generated semantics.}, + isbn = {978-3-030-45234-6} +} + +@InProceedings{ Mavridou2019, + author = {Mavridou, Anastasia and Laszka, Aron and Stachtiari, Emmanouela and Dubey, Abhishek}, + booktitle = {Financial Cryptography and Data Security}, + title = {VeriSolid: Correct-by-Design Smart Contracts for Ethereum}, + year = 2019, + address = {Cham}, + editor = {Goldberg, Ian and Moore, Tyler}, + pages = {446--465}, + publisher = {Springer International Publishing}, + abstract = {The adoption of blockchain based distributed ledgers is growing fast due to their ability to provide + reliability, integrity, and auditability without trusted entities. One of the key capabilities of + these emerging platforms is the ability to create self-enforcing smart contracts. However, the + development of smart contracts has proven to be error-prone in practice, and as a result, contracts + deployed on public platforms are often riddled with security vulnerabilities. This issue is + exacerbated by the design of these platforms, which forbids updating contract code and rolling back + malicious transactions. In light of this, it is crucial to ensure that a smart contract is secure + before deploying it and trusting it with significant amounts of cryptocurrency. To this end, we + introduce the VeriSolid framework for the formal verification of contracts that are specified using a + transition-system based model with rigorous operational semantics. Our model-based approach allows + developers to reason about and verify contract behavior at a high level of abstraction. VeriSolid + allows the generation of Solidity code from the verified models, which enables the correct-by-design + development of smart{\^A}contracts.}, + isbn = {978-3-030-32101-7} +} + +@InProceedings{ Mavridou2018, + author = {Mavridou, Anastasia and Laszka, Aron}, + booktitle = {Principles of Security and Trust}, + title = {Tool Demonstration: FSolidM for Designing Secure Ethereum Smart Contracts}, + year = 2018, + address = {Cham}, + editor = {Bauer, Lujo and K{\"u}sters, Ralf}, + pages = {270--277}, + publisher = {Springer International Publishing}, + abstract = {Blockchain-based distributed computing platforms enable the trusted execution of + computation---defined in the form of smart contracts---without trusted agents. Smart contracts are + envisioned to have a variety of applications, ranging from financial to IoT asset tracking. + Unfortunately, the development of smart contracts has proven to be extremely error prone. In practice, + contracts are riddled with security vulnerabilities comprising a critical issue since bugs are by + design non-fixable and contracts may handle financial assets of significant value. To facilitate the + development of secure smart contracts, we have created the FSolidM framework, which allows developers + to define contracts as finite state machines (FSMs) with rigorous and clear semantics. FSolidM + provides an easy-to-use graphical editor for specifying FSMs, a code generator for creating Ethereum + smart contracts, and a set of plugins that developers may add to their FSMs to enhance security and + functionality.}, + isbn = {978-3-319-89722-6} +} + +@InProceedings{ Ahrendt2020, + author = {Ahrendt, Wolfgang and Bubel, Richard}, + booktitle = {Leveraging Applications of Formal Methods, Verification and Validation: Applications}, + title = {Functional Verification of Smart Contracts via Strong Data Integrity}, + year = 2020, + address = {Cham}, + editor = {Margaria, Tiziana and Steffen, Bernhard}, + pages = {9--24}, + publisher = {Springer International Publishing}, + abstract = {We present an invariant-based specification and verification methodology that allows us to + conveniently specify and verify strong data integrity properties for Solidity smart contracts. Our + approach is able to reason precisely about arbitrary usage of the contracts, which may include + re-entrance, a common security pitfall in smart contracts. We implemented the approach in a prototype + verification tool, called SolidiKeY, and applied it successfully to a number of smart contracts.}, + isbn = {978-3-030-61467-6} +} + +@Article{ Ahrendt2016, + author = {Ahrendt, Wolfgang and Beckert, Bernhard and Bubel, Richard and H{\"a}hnle, Reiner and Schmitt, Peter + H and Ulbrich, Mattias}, + journal = {Lecture Notes in Computer Science}, + title = {Deductive software verification--the key book}, + year = 2016, + volume = 10001, + publisher = {Springer} +} + +@InProceedings{ Crafa2020, + author = {Crafa, Silvia and Di Pirro, Matteo and Zucca, Elena}, + booktitle = {Financial Cryptography and Data Security}, + title = {Is Solidity Solid Enough?}, + year = 2020, + address = {Cham}, + editor = {Bracciali, Andrea and Clark, Jeremy and Pintore, Federico and R{\o}nne, Peter B. and Sala, + Massimiliano}, + pages = {138--153}, + publisher = {Springer International Publishing}, + abstract = {We introduce Featherweight Solidity, a calculus formalizing the core features of the Solidity + language, thus providing a fundamental step to reason about safety properties of smart contracts' + source code. The formalization includes a static type system that represents the foundation of the + Solidity compiler. We show that it prevents some errors whereas many others, such as accesses to a non + existing function or state variable, are only detected at runtime and cause interruption and + rolling-back of transactions. We then propose a refinement of the type system that is retro-compatible + with original Solidity code, and statically captures more errors, such as unsafe casts and unsafe + call-back expressions.}, + isbn = {978-3-030-43725-1} +} + +@Misc{ Remix, + author = {Online}, + howpublished = {\url{https://remix-ide.readthedocs.io/en/latest}}, + title = {Remix - Solidity IDE} +} + +@InProceedings{ Bartoletti2019, + author = {Bartoletti, Massimo and Galletta, Letterio and Murgia, Maurizio}, + booktitle = {Data Privacy Management, Cryptocurrencies and Blockchain Technology}, + title = {A Minimal Core Calculus for Solidity Contracts}, + year = 2019, + address = {Cham}, + editor = {P{\'e}rez-Sol{\`a}, Cristina and Navarro-Arribas, Guillermo and Biryukov, Alex and Garcia-Alfaro, + Joaquin}, + pages = {233--243}, + publisher = {Springer International Publishing}, + abstract = {The Ethereum platform supports the decentralized execution of smart contracts, i.e. computer programs + that transfer digital assets between users. The most common language used to develop these contracts + is Solidity, a Javascript-like language which compiles into EVM bytecode, the language actually + executed by Ethereum nodes. While much research has addressed the formalisation of the semantics of + EVM bytecode, relatively little attention has been devoted to that of Solidity. In this paper we + propose a minimal calculus for Solidity contracts, which extends an imperative core with a single + primitive to transfer currency and invoke contract procedures. We build upon this formalisation to + give semantics to the Ethereum blockchain. We show our calculus expressive enough to reason about some + typical quirks of Solidity, like e.g. re-entrancy.}, + isbn = {978-3-030-31500-9} +} + +@InProceedings{ Zakrzewski2018, + author = {Jakub Zakrzewski}, + booktitle = {Verified Software. Theories, Tools, and Experiments - 10th International Conference, {VSTTE} 2018, + Oxford, UK, July 18-19, 2018, Revised Selected Papers}, + title = {Towards Verification of Ethereum Smart Contracts: {A} Formalization of Core of Solidity}, + year = 2018, + editor = {Ruzica Piskac and Philipp R{\"{u}}mmer}, + pages = {229--247}, + publisher = {Springer}, + series = {Lecture Notes in Computer Science}, + volume = 11294, + bibsource = {dblp computer science bibliography, https://dblp.org}, + biburl = {https://dblp.org/rec/conf/vstte/Zakrzewski18.bib}, + doi = {10.1007/978-3-030-03592-1\_13}, + timestamp = {Tue, 14 May 2019 10:00:49 +0200}, + url = {https://doi.org/10.1007/978-3-030-03592-1\_13} +} + +@InProceedings{ Hajdu2020, + author = {Hajdu, {\'A}kos and Jovanovi{\'{c}}, Dejan}, + booktitle = {Verified Software. Theories, Tools, and Experiments}, + title = {solc-verify: A Modular Verifier for Solidity Smart Contracts}, + year = 2020, + address = {Cham}, + editor = {Chakraborty, Supratik and Navas, Jorge A.}, + pages = {161--179}, + publisher = {Springer International Publishing}, + abstract = {We present solc-verify, a source-level verification tool for Ethereum smart contracts. solc-verify + takes smart contracts written in Solidity and discharges verification conditions using modular program + analysis and SMT solvers. Built on top of the Solidity compiler, solc-verify reasons at the level of + the contract source code, as opposed to the more common approaches that operate at the level of + Ethereum bytecode. This enables solc-verify to effectively reason about high-level contract properties + while modeling low-level language semantics precisely. The properties, such as contract invariants, + loop invariants, and function pre- and post-conditions, can be provided as annotations in the code by + the developer. This enables automated, yet user-friendly formal verification for smart contracts. We + demonstrate solc-verify by examining real-world examples where our tool can effectively find bugs and + prove correctness of non-trivial properties with minimal user effort.}, + isbn = {978-3-030-41600-3} +} + +@InProceedings{ Amani2018, + author = {Amani, Sidney and B\'{e}gel, Myriam and Bortin, Maksym and Staples, Mark}, + booktitle = {Proceedings of the 7th ACM SIGPLAN International Conference on Certified Programs and Proofs}, + title = {Towards Verifying Ethereum Smart Contract Bytecode in Isabelle/HOL}, + year = 2018, + address = {New York, NY, USA}, + pages = {66{\^a}77}, + publisher = {Association for Computing Machinery}, + series = {CPP 2018}, + abstract = {Blockchain technology has increasing attention in research and across many industries. The Ethereum + blockchain offers smart contracts, which are small programs defined, executed, and recorded as + transactions in the blockchain transaction history. These smart contracts run on the Ethereum Virtual + Machine (EVM) and can be used to encode agreements, transfer assets, and enforce integrity conditions + in relationships between parties. Smart contracts can carry financial value, and are increasingly used + for safety-, security-, or mission-critical purposes. Errors in smart contracts have led and will lead + to loss or harm. Formal verification can provide the highest level of confidence about the correct + behaviour of smart contracts. In this paper we extend an existing EVM formalisation in Isabelle/HOL by + a sound program logic at the level of bytecode. We structure bytecode sequences into blocks of + straight-line code and create a program logic to reason about these. This abstraction is a step + towards control of the cost and complexity of formal verification of EVM smart contracts.}, + doi = {10.1145/3167084}, + isbn = 9781450355865, + keywords = {blockchain, formal verification, smart contracts, Ethereum, Isabelle/HOL}, + location = {Los Angeles, CA, USA}, + numpages = 12, + url = {https://doi.org/10.1145/3167084} +} + +@Misc{ Wood2014, + author = {Wood, Gavin and others}, + journal = {Ethereum project yellow paper}, + title = {Ethereum: A secure decentralised generalised transaction ledger}, + year = 2022, + note = {Berlin Version 3078285 -- 2022-07-13. \url{https://ethereum.github.io/yellowpaper/paper.pdf}}, + pages = {1--32}, +} + +@Misc{ Solidity, + author = {Online}, + howpublished = {\url{https://solidity.readthedocs.io/en/latest}}, + title = {Solidity documentation} +} + +@InProceedings{ Grishchenko2018, + author = {Grishchenko, Ilya and Maffei, Matteo and Schneidewind, Clara}, + booktitle = {Principles of Security and Trust}, + title = {A Semantic Framework for the Security Analysis of Ethereum Smart Contracts}, + year = 2018, + address = {Cham}, + editor = {Bauer, Lujo and K{\"u}sters, Ralf}, + pages = {243--269}, + publisher = {Springer International Publishing}, + abstract = {Smart contracts are programs running on cryptocurrency (e.g., Ethereum) blockchains, whose popularity + stem from the possibility to perform financial transactions, such as payments and auctions, in a + distributed environment without need for any trusted third party. Given their financial nature, bugs + or vulnerabilities in these programs may lead to catastrophic consequences, as witnessed by recent + attacks. Unfortunately, programming smart contracts is a delicate task that requires strong expertise: + Ethereum smart contracts are written in Solidity, a dedicated language resembling JavaScript, and + shipped over the blockchain in the EVM bytecode format. In order to rigorously verify the security of + smart contracts, it is of paramount importance to formalize their semantics as well as the security + properties of interest, in particular at the level of the bytecode being executed.}, + isbn = {978-3-319-89722-6} +} + +@Article{ Wood2014a, + author = {Wood, Gavin and others}, + journal = {Ethereum project yellow paper}, + title = {Ethereum: A secure decentralised generalised transaction ledger}, + year = 2014, + number = 2014, + pages = {1--32}, + volume = 151 +} + +@InProceedings{ Swamy2016, + author = {Swamy, Nikhil and Hri\c{t}cu, C\u{a}t\u{a}lin and Keller, Chantal and Rastogi, Aseem and + Delignat-Lavaud, Antoine and Forest, Simon and Bhargavan, Karthikeyan and Fournet, C\'{e}dric and + Strub, Pierre-Yves and Kohlweiss, Markulf and Zinzindohoue, Jean-Karim and Zanella-B\'{e}guelin, Santiago}, + booktitle = {Proceedings of the 43rd Annual ACM SIGPLAN-SIGACT Symposium on Principles of Programming Languages}, + title = {Dependent Types and Multi-Monadic Effects in F*}, + year = 2016, + address = {New York, NY, USA}, + pages = {256{\^a}270}, + publisher = {Association for Computing Machinery}, + series = {POPL '16}, + abstract = {We present a new, completely redesigned, version of F*, a language that works both as a proof + assistant as well as a general-purpose, verification-oriented, effectful programming language. In + support of these complementary roles, F* is a dependently typed, higher-order, call-by-value language + with _primitive_ effects including state, exceptions, divergence and IO. Although primitive, + programmers choose the granularity at which to specify effects by equipping each effect with a + monadic, predicate transformer semantics. F* uses this to efficiently compute weakest preconditions + and discharges the resulting proof obligations using a combination of SMT solving and manual proofs. + Isolated from the effects, the core of F* is a language of pure functions used to write specifications + and proof terms---its consistency is maintained by a semantic termination check based on a + well-founded order. We evaluate our design on more than 55,000 lines of F* we have authored in the + last year, focusing on three main case studies. Showcasing its use as a general-purpose programming + language, F* is programmed (but not verified) in F*, and bootstraps in both OCaml and F#. Our + experience confirms F*'s pay-as-you-go cost model: writing idiomatic ML-like code with no finer + specifications imposes no user burden. As a verification-oriented language, our most significant + evaluation of F* is in verifying several key modules in an implementation of the TLS-1.2 protocol + standard. For the modules we considered, we are able to prove more properties, with fewer annotations + using F* than in a prior verified implementation of TLS-1.2. Finally, as a proof assistant, we discuss + our use of F* in mechanizing the metatheory of a range of lambda calculi, starting from the simply + typed lambda calculus to System F-omega and even micro-F*, a sizeable fragment of F* itself---these + proofs make essential use of F*'s flexible combination of SMT automation and constructive proofs, + enabling a tactic-free style of programming and proving at a relatively large scale.}, + doi = {10.1145/2837614.2837655}, + isbn = 9781450335492, + keywords = {effectful programming, verification, proof assistants}, + location = {St. Petersburg, FL, USA}, + numpages = 15, + url = {https://doi.org/10.1145/2837614.2837655} +} + +@Article{ 10.1145/2914770.2837655, + author = {Swamy, Nikhil and Hri\c{t}cu, C\u{a}t\u{a}lin and Keller, Chantal and Rastogi, Aseem and + Delignat-Lavaud, Antoine and Forest, Simon and Bhargavan, Karthikeyan and Fournet, C\'{e}dric and + Strub, Pierre-Yves and Kohlweiss, Markulf and Zinzindohoue, Jean-Karim and Zanella-B\'{e}guelin, Santiago}, + journal = {SIGPLAN Not.}, + title = {Dependent Types and Multi-Monadic Effects in F*}, + year = 2016, + issn = {0362-1340}, + month = jan, + number = 1, + pages = {256{\^a}270}, + volume = 51, + abstract = {We present a new, completely redesigned, version of F*, a language that works both as a proof + assistant as well as a general-purpose, verification-oriented, effectful programming language. In + support of these complementary roles, F* is a dependently typed, higher-order, call-by-value language + with _primitive_ effects including state, exceptions, divergence and IO. Although primitive, + programmers choose the granularity at which to specify effects by equipping each effect with a + monadic, predicate transformer semantics. F* uses this to efficiently compute weakest preconditions + and discharges the resulting proof obligations using a combination of SMT solving and manual proofs. + Isolated from the effects, the core of F* is a language of pure functions used to write specifications + and proof terms---its consistency is maintained by a semantic termination check based on a + well-founded order. We evaluate our design on more than 55,000 lines of F* we have authored in the + last year, focusing on three main case studies. Showcasing its use as a general-purpose programming + language, F* is programmed (but not verified) in F*, and bootstraps in both OCaml and F#. Our + experience confirms F*'s pay-as-you-go cost model: writing idiomatic ML-like code with no finer + specifications imposes no user burden. As a verification-oriented language, our most significant + evaluation of F* is in verifying several key modules in an implementation of the TLS-1.2 protocol + standard. For the modules we considered, we are able to prove more properties, with fewer annotations + using F* than in a prior verified implementation of TLS-1.2. Finally, as a proof assistant, we discuss + our use of F* in mechanizing the metatheory of a range of lambda calculi, starting from the simply + typed lambda calculus to System F-omega and even micro-F*, a sizeable fragment of F* itself---these + proofs make essential use of F*'s flexible combination of SMT automation and constructive proofs, + enabling a tactic-free style of programming and proving at a relatively large scale.}, + address = {New York, NY, USA}, + doi = {10.1145/2914770.2837655}, + issue_date = {January 2016}, + keywords = {proof assistants, verification, effectful programming}, + numpages = 15, + publisher = {Association for Computing Machinery}, + url = {https://doi.org/10.1145/2914770.2837655} +} + +@Article{ Mulligan2014, + author = {Mulligan, Dominic P. and Owens, Scott and Gray, Kathryn E. and Ridge, Tom and Sewell, Peter}, + journal = {SIGPLAN Not.}, + title = {Lem: Reusable Engineering of Real-World Semantics}, + year = 2014, + issn = {0362-1340}, + month = aug, + number = 9, + pages = {175{\^a}188}, + volume = 49, + abstract = {Recent years have seen remarkable successes in rigorous engineering: using mathematically rigorous + semantic models (not just idealised calculi) of real-world processors, programming languages, + protocols, and security mechanisms, for testing, proof, analysis, and design. Building these models is + challenging, requiring experimentation, dialogue with vendors or standards bodies, and validation; + their scale adds engineering issues akin to those of programming to the task of writing clear and + usable mathematics. But language and tool support for specification is lacking. Proof assistants can + be used but bring their own difficulties, and a model produced in one, perhaps requiring many + person-years effort and maintained over an extended period, cannot be used by those familiar with + another.We introduce Lem, a language for engineering reusable large-scale semantic models. The Lem + design takes inspiration both from functional programming languages and from proof assistants, and Lem + definitions are translatable into OCaml for testing, Coq, HOL4, and Isabelle/HOL for proof, and LaTeX + and HTML for presentation. This requires a delicate balance of expressiveness, careful library design, + and implementation of transformations - akin to compilation, but subject to the constraint of + producing usable and human-readable code for each target. Lem's effectiveness is demonstrated by its + use in practice.}, + address = {New York, NY, USA}, + doi = {10.1145/2692915.2628143}, + issue_date = {September 2014}, + keywords = {proof assistants, real-world semantics, lem, specification languages}, + numpages = 14, + publisher = {Association for Computing Machinery}, + url = {https://doi.org/10.1145/2692915.2628143} +} + +@InProceedings{ 10.1145/2628136.2628143, + author = {Mulligan, Dominic P. and Owens, Scott and Gray, Kathryn E. and Ridge, Tom and Sewell, Peter}, + booktitle = {Proceedings of the 19th ACM SIGPLAN International Conference on Functional Programming}, + title = {Lem: Reusable Engineering of Real-World Semantics}, + year = 2014, + address = {New York, NY, USA}, + pages = {175{\^a}188}, + publisher = {Association for Computing Machinery}, + series = {ICFP '14}, + abstract = {Recent years have seen remarkable successes in rigorous engineering: using mathematically rigorous + semantic models (not just idealised calculi) of real-world processors, programming languages, + protocols, and security mechanisms, for testing, proof, analysis, and design. Building these models is + challenging, requiring experimentation, dialogue with vendors or standards bodies, and validation; + their scale adds engineering issues akin to those of programming to the task of writing clear and + usable mathematics. But language and tool support for specification is lacking. Proof assistants can + be used but bring their own difficulties, and a model produced in one, perhaps requiring many + person-years effort and maintained over an extended period, cannot be used by those familiar with + another.We introduce Lem, a language for engineering reusable large-scale semantic models. The Lem + design takes inspiration both from functional programming languages and from proof assistants, and Lem + definitions are translatable into OCaml for testing, Coq, HOL4, and Isabelle/HOL for proof, and LaTeX + and HTML for presentation. This requires a delicate balance of expressiveness, careful library design, + and implementation of transformations - akin to compilation, but subject to the constraint of + producing usable and human-readable code for each target. Lem's effectiveness is demonstrated by its + use in practice.}, + doi = {10.1145/2628136.2628143}, + isbn = 9781450328739, + keywords = {real-world semantics, lem, specification languages, proof assistants}, + location = {Gothenburg, Sweden}, + numpages = 14, + url = {https://doi.org/10.1145/2628136.2628143} +} + +@Manual{ solidity-0.8.4, + title = {Solidity (Version 0.8.4)}, + howpublished = {\url{https://docs.soliditylang.org/en/v0.8.4/}} +} + +@TechReport{ wood:ethereum:2021, + author = {Gavin Wood}, + title = {Ethereum: A Secure Decentralised Generalised Transation Ledger (Version 2021-04-21)}, + optnote = {\url{https://ethereum.github.io/yellowpaper/paper.pdf}} +} + +@Misc{ coinmarket, + title = {The {Bitcon} Market Capitalisation.}, + url = {https://coinmarketcap.com/currencies/bitcoin/}, + note = {Last checked on 2021-05-04.} +} + +@InCollection{ marmsoler.ea:conformance:2022, + abstract = {A common problem in verification is to ensure that the formal specification models the real-world + system, i.e., the implementation, faithfully. Testing is a technique that can help to bridge the gap + between a formal specification and its implementation.\\\\Fuzzing in general and grammar-based fuzzing + in particular are successfully used for finding bugs in implementations. Traditional fuzzing + applications rely on an implicit test specification that informally can be described as ``the program + under test does not crash''.\\\\In this paper, we present an approach using grammar-based fuzzing to + ensure the conformance of a formal specification, namely the formal semantics of the Solidity + Programming language, to a real-world implementation. For this, we derive an executable test-oracle + from the formal semantics of Solidity in Isabelle/HOL. The derived test oracle is used during the + fuzzing of the implementation to validate that the formal semantics and the implementation are in + conformance.}, + address = {Heidelberg}, + author = {Diego Marmsoler and Achim D. Brucker}, + booktitle = {{TAP} 2022: Tests And Proofs}, + editor = {Laura Kovacs and Karl Meinke}, + isbn = {978-3-642-38915-3}, + keywords = {Conformance Testing, Fuzzing, Verification, Solidity}, + language = {USenglish}, + location = {Nantes, France}, + pdf = {https://www.brucker.ch/bibliography/download/2022/marmsoler.ea-conformance-2022.pdf}, + publisher = {Springer-Verlag}, + series = {Lecture Notes in Computer Science}, + title = {Conformance Testing of Formal Semantics using Grammar-based Fuzzing}, + url = {https://www.brucker.ch/bibliography/abstract/marmsoler.ea-conformance-2022}, + year = 2022 +} + +@InCollection{ marmsoler.ea:solidity-semantics:2021, + abstract = {Smart contracts are programs, usually automating legal agreements such as financial transactions. + Thus, bugs in smart contracts can lead to large financial losses. For example, an incorrectly + initialized contract was the root cause of the Parity Wallet bug that made USD 280mil worth of Ether + inaccessible. Ether is the cryptocurrency of the Ethereum blockchain that uses Solidity for expressing + smart contracts.\\\\In this paper, we address this problem by presenting an executable denotational + semantics for Solidity in the interactive theorem prover Isabelle/HOL. This formal semantics builds + the foundation of an interactive program verification environment for Solidity programs and allows for + inspecting Solidity programs by (symbolic) execution. We combine the latter with grammar-based fuzzing + to ensure that our formal semantics complies to the Solidity implementation on the Ethereum + Blockchain. Finally, we demonstrate the formal verification of Solidity programs by two examples: + constant folding and memory optimization.}, + address = {Heidelberg}, + author = {Diego Marmsoler and Achim D. Brucker}, + booktitle = {Software Engineering and Formal Methods (SEFM)}, + editor = {Radu Calinescu and Corina Pasareanu}, + isbn = {3-540-25109-X}, + keywords = {Solidity, Denotational Semantics, Isabelle/HOL, Gas Optimization}, + language = {USenglish}, + pdf = {https://www.brucker.ch/bibliography/download/2021/marmsoler.ea-solidity-semantics-2021.pdf}, + publisher = {Springer-Verlag}, + series = {Lecture Notes in Computer Science}, + title = {A Denotational Semantics of {Solidity} in {Isabelle/HOL}}, + url = {https://www.brucker.ch/bibliography/abstract/marmsoler.ea-solidity-semantics-2021}, + year = 2021 +} +@COMMENT{jabref-meta: databaseType:bibtex;} diff --git a/thys/Solidity/document/root.tex b/thys/Solidity/document/root.tex --- a/thys/Solidity/document/root.tex +++ b/thys/Solidity/document/root.tex @@ -1,228 +1,236 @@ -\RequirePackage{ifvtex} -\documentclass[10pt,DIV17,a4paper,abstract=true,twoside=semi,openright] -{scrreprt} -\usepackage[T1]{fontenc} -\usepackage[utf8]{inputenc} -\usepackage{textcomp} -\bibliographystyle{abbrvnat} -\usepackage[english]{babel} -\RequirePackage[caption]{subfig} -\usepackage{isabelle} -\usepackage{isabellesym} -\usepackage{amsmath} -\usepackage{amssymb} -\usepackage[numbers, sort&compress, sectionbib]{natbib} -\usepackage{graphicx} -\usepackage{hyperref} -\usepackage{orcidlink} -\setcounter{tocdepth}{3} -\hypersetup{% - bookmarksdepth=3 - ,pdfpagelabels - ,pageanchor=true - ,bookmarksnumbered - ,plainpages=false -} % more detailed digital TOC (aka bookmarks) -\sloppy -\allowdisplaybreaks[4] -\urlstyle{rm} -\isabellestyle{it} - -% for uniform font size -%\renewcommand{\isastyle}{\isastyleminor} - -\newenvironment{frontmatter}{}{} - -\pagestyle{headings} -\isabellestyle{default} -\setcounter{tocdepth}{1} -\newcommand{\ie}{i.\,e.\xspace} -\newcommand{\eg}{e.\,g.\xspace} -\newcommand{\thy}{\isabellecontext} -\renewcommand{\isamarkupsection}[1]{% - \begingroup% - \def\isacharunderscore{\textunderscore}% - \section{#1 (\thy)}% - \endgroup% -} - - -\title{Isabelle/Solidity} -\subtitle{A deep Embedding of Solidity in Isabelle/HOL} -\author{Diego Marmsoler\textsuperscript{\orcidlink{0000-0003-2859-7673}} - and - Achim D. Brucker\textsuperscript{\orcidlink{0000-0002-6355-1200}}}% -\publishers{ -Department of Computer Science, University of Exeter, Exeter, UK\texorpdfstring{\\}{, } - \texttt{\{d.marmsoler,a.brucker\}@exeter.ac.uk} -} -\begin{document} -\begin{frontmatter} -\maketitle -\begin{abstract} - \begin{quote} - Smart contracts are automatically executed programs, usually representing - legal agreements such as financial transactions. Thus, bugs in smart - contracts can lead to large financial losses. For example, an incorrectly - initialized contract was the root cause of the Parity Wallet bug that saw - \$280M worth of Ether destroyed. Ether is the cryptocurrency of the - Ethereum blockchain that uses Solidity for expressing smart contracts. - - We address this problem by formalizing an executable denotational semantics - for Solidity in the interactive theorem prover Isabelle/HOL. This formal - semantics builds the foundation of an interactive program verification - environment for Solidity programs and allows for inspecting them by - (symbolic) execution. We combine the latter with grammar based fuzzing to - ensure that our formal semantics complies to the Solidity implementation on - the Ethereum Blockchain. Finally, we demonstrate the formal verification of - Solidity programs by two examples: constant folding and a simple verified - token. - - \bigskip - \noindent\textbf{Keywords:} {Solidity, Denotational Semantics, - Isabelle/HOL, Gas} - \end{quote} -\end{abstract} - -\tableofcontents -\cleardoublepage -\end{frontmatter} - - -\chapter{Introduction} -An increasing number of businesses is adopting blockchain-based solutions. Most -notably, the market value of Bitcoin, most likely the first and most well-known -blockchain-based cryptocurrency, passed USD 1 trillion in February -2021~\cite{coinmarket}. While Bitcoin might be the most well-known application -of a blockchain, it lacks features that applications outside cryptocurrencies -require and that make blockchain solutions attractive to businesses. - -For example, the Ethereum blockchain~\cite{Wood2014} is a feature-rich -distributed computing platform that provides not only a cryptocurrency, called -\emph{Ether}: Ethereum also provides an immutable distributed data structure -(the \emph{blockchain}) on which distributed programs, called \emph{smart -contracts}, can be executed. Essentially, smart contracts are automatically -executed programs, usually representing a legal agreement, e.g., financial -transactions. To support those applications, Ethereum provides a dedicated -account data structure on its blockchain that smart contracts can modify, i.e., -transferring Ether between accounts. Thus, bugs in smart contracts can lead to -large financial losses. For example, an incorrectly initialized contract was -the root cause of the Parity Wallet bug that saw \$280M worth of Ether -destroyed~\cite{perez.ea:smart:2021}. This risk of bugs being costly is already -a big motivation for using formal verification techniques to minimize this risk. -The fact that smart contracts are deployed on the blockchain immutably, i.e., -they cannot be updated or removed easily, makes it even more important to ``get -smart contracts'' right, before they are deployed on a blockchain for the very -first time. - -For implementing smart contracts, Ethereum provides -\emph{Solidity}~\cite{Solidity}, a Turing-complete, statically typed programming -language that has been designed to look familiar to people knowing Java, C, or -JavaScript. Notably, the type system provides, e.g., numerous integer types of -different sizes (e.g., \texttt{uint256}) and Solidity also relies on different -types of stores. While Solidity is Turing-complete, the execution of Solidity -programs is guaranteed to terminate. The reason for this is that executing -Solidity operations costs \emph{gas}, a tradable commodity on the Ethereum -blockchain. Gas does cost Ether and hence, programmers of smart contracts have -an incentive to write highly optimized contracts whose execution consumes as -little gas as possible. For example, the size of the integer types used can -impact the amount of gas required for executing a contract. This desire for -highly optimized contracts can conflict with the desire to write correct -contracts. - -In this paper, we address the problem of developing smart contracts in Solidity -that are correct: we present an executable denotational semantics for Solidity -in the interactive theorem prover Isabelle/HOL. - -In particular, our semantics supports the following features of Solidity: -\begin{itemize} - \item \emph{Fixed-size integer types} of various lengths and corresponding arithmetic. - \item \emph{Domain-specific primitives}, such as money transfer or balance queries. - \item \emph{Different types of stores}, such as storage, memory, and stack. - \item \emph{Complex data types}, such as hash-maps and arrays. - \item \emph{Assignments with different semantics}, depending on data types. - \item An extendable \emph{gas model}. - \item \emph{Internal and external method calls}. -\end{itemize} - -A more abstract description of the semantics is given -in~\cite{marmsoler.ea:solidity-semantics:2021} and the conformance testing -approach for ensuring that our semantics conforms to the actual -implementation is described in~\cite{marmsoler.ea:conformance:2022}. - - -The rest of this document is automatically generated from the formalization in -Isabelle/HOL, i.e., all content is checked by Isabelle. The structure follows -the theory dependencies (see \autoref{fig:session-graph}). - -\begin{figure} - \centering - \includegraphics[height=.9\textheight, width=\textwidth,keepaspectratio]{session_graph} - \caption{The Dependency Graph of the Isabelle Theories.\label{fig:session-graph}} -\end{figure} - -\clearpage -%\input{session} - -\chapter{Preliminaries} -In this chapter, we discuss auxiliary formalizations and functions that -are used in our Solidity semantics but are more generic, i.e., not -specific to Solidity. This includes, for example, functions to convert -values of basic types to/from strings. - -\input{ReadShow.tex} - -\input{StateMonad.tex} - - -\chapter{Types and Accounts} -In this chapter, we discuss the basic data types of Solidity and the -representations of accounts. - -\input{Valuetypes.tex} - -\chapter{Stores and Environment} -In this chapter, we focus on a particular aspect of Solidity that is -different to most programming languages: the handling of memory in -general and, in particular, the different between store and storage. - -\input{Storage.tex} -\input{Environment.tex} - -\chapter{Expressions and Statements} - -In this chapter, we formalize expressions, declarations, and -statements. The results up to here form the core of our Solidity -semantics. -\input{Statements.tex} -\input{Solidity_Main.tex} - - -\chapter{A Solidity Evaluation System} -This chapter discussed a tactic for symbolically executing Solidity statements -and expressions as well as provides a configuration for Isabelle's code -generator that allows us to generate an efficient implementation of our -executable formal semantics in, e.g., Haskell, SML, or Scala. In our test -framework, we use Haskell as a target language. -\input{Solidity_Symbex.tex} -\input{Solidity_Evaluator.tex} -\IfFileExists{Compile_Evaluator.tex}{% - \input{Compile_Evaluator.tex} -}{} - -\chapter{Applications} -In this chapter, we discuss various applications of our Solidity -semantics. -\input{Constant_Folding.tex} -\input{Reentrancy.tex} - - -\IfFileExists{root.bib}{% - \bibliography{root} -}{} -\end{document} - -%%% Local Variables: -%%% mode: latex -%%% TeX-master: t -%%% End: +\RequirePackage{ifvtex} +\documentclass[10pt,DIV17,a4paper,abstract=true,twoside=semi,openright] +{scrreprt} +\usepackage[T1]{fontenc} +\usepackage[utf8]{inputenc} +\usepackage{textcomp} +\bibliographystyle{abbrvnat} +\usepackage[english]{babel} +\RequirePackage[caption]{subfig} +\usepackage{isabelle} +\usepackage{isabellesym} +\usepackage{amsmath} +\usepackage{amssymb} +\usepackage[numbers, sort&compress, sectionbib]{natbib} +\usepackage{graphicx} +\usepackage{hyperref} +\usepackage{orcidlink} +\setcounter{tocdepth}{3} +\hypersetup{% + bookmarksdepth=3 + ,pdfpagelabels + ,pageanchor=true + ,bookmarksnumbered + ,plainpages=false +} % more detailed digital TOC (aka bookmarks) +\sloppy +\allowdisplaybreaks[4] +\urlstyle{rm} +\isabellestyle{it} + +% for uniform font size +%\renewcommand{\isastyle}{\isastyleminor} + +\newenvironment{frontmatter}{}{} + +\pagestyle{headings} +\isabellestyle{default} +\setcounter{tocdepth}{1} +\newcommand{\ie}{i.\,e.\xspace} +\newcommand{\eg}{e.\,g.\xspace} +\newcommand{\thy}{\isabellecontext} +\renewcommand{\isamarkupsection}[1]{% + \begingroup% + \def\isacharunderscore{\textunderscore}% + \section{#1 (\thy)}% + \endgroup% +} + + +\title{Isabelle/Solidity} +\subtitle{A deep Embedding of Solidity in Isabelle/HOL} +\author{Diego Marmsoler\textsuperscript{\orcidlink{0000-0003-2859-7673}} + and + Achim D. Brucker\textsuperscript{\orcidlink{0000-0002-6355-1200}} + and Billy Thornton}% +\publishers{ +Department of Computer Science, University of Exeter, Exeter, UK\texorpdfstring{\\}{, } + \texttt{\{d.marmsoler, a.brucker, bt319\}@exeter.ac.uk} +} +\begin{document} +\begin{frontmatter} +\maketitle +\begin{abstract} + \begin{quote} + Smart contracts are automatically executed programs, usually representing + legal agreements such as financial transactions. Thus, bugs in smart + contracts can lead to large financial losses. For example, an incorrectly + initialized contract was the root cause of the Parity Wallet bug that saw + \$280M worth of Ether destroyed. Ether is the cryptocurrency of the + Ethereum blockchain that uses Solidity for expressing smart contracts. + + We address this problem by formalizing an executable denotational semantics + for Solidity in the interactive theorem prover Isabelle/HOL. This formal + semantics builds the foundation of an interactive program verification + environment for Solidity programs and allows for inspecting them by + (symbolic) execution. We combine the latter with grammar based fuzzing to + ensure that our formal semantics complies to the Solidity implementation on + the Ethereum blockchain. Finally, we demonstrate the formal verification of + Solidity programs by two examples: constant folding and a simple verified + token. + + \bigskip + \noindent\textbf{Keywords:} {Solidity, Denotational Semantics, + Isabelle/HOL, Gas} + \end{quote} +\end{abstract} + +\tableofcontents +\cleardoublepage +\end{frontmatter} + + +\chapter{Introduction} +An increasing number of businesses is adopting blockchain-based solutions. Most +notably, the market value of Bitcoin, most likely the first and most well-known +blockchain-based cryptocurrency, passed USD 1 trillion in February +2021~\cite{coinmarket}. While Bitcoin might be the most well-known application +of a blockchain, it lacks features that applications outside cryptocurrencies +require and that make blockchain solutions attractive to businesses. + +For example, the Ethereum blockchain~\cite{Wood2014} is a feature-rich +distributed computing platform that provides not only a cryptocurrency, called +\emph{Ether}: Ethereum also provides an immutable distributed data structure +(the \emph{blockchain}) on which distributed programs, called \emph{smart +contracts}, can be executed. Essentially, smart contracts are automatically +executed programs, usually representing a legal agreement, e.g., financial +transactions. To support those applications, Ethereum provides a dedicated +account data structure on its blockchain that smart contracts can modify, i.e., +transferring Ether between accounts. Thus, bugs in smart contracts can lead to +large financial losses. For example, an incorrectly initialized contract was +the root cause of the Parity Wallet bug that saw \$280M worth of Ether +destroyed~\cite{perez.ea:smart:2021}. This risk of bugs being costly is already +a big motivation for using formal verification techniques to minimize this risk. +The fact that smart contracts are deployed on the blockchain immutably, i.e., +they cannot be updated or removed easily, makes it even more important to ``get +smart contracts'' right, before they are deployed on a blockchain for the very +first time. + +For implementing smart contracts, Ethereum provides +\emph{Solidity}~\cite{Solidity}, a Turing-complete, statically typed programming +language that has been designed to look familiar to people knowing Java, C, or +JavaScript. Notably, the type system provides, e.g., numerous integer types of +different sizes (e.g., \texttt{uint256}) and Solidity also relies on different +types of stores. While Solidity is Turing-complete, the execution of Solidity +programs is guaranteed to terminate. The reason for this is that executing +Solidity operations costs \emph{gas}, a tradable commodity on the Ethereum +blockchain. Gas does cost Ether and hence, programmers of smart contracts have +an incentive to write highly optimized contracts whose execution consumes as +little gas as possible. For example, the size of the integer types used can +impact the amount of gas required for executing a contract. This desire for +highly optimized contracts can conflict with the desire to write correct +contracts. + +In this paper, we address the problem of developing smart contracts in Solidity +that are correct: we present an executable denotational semantics for Solidity +in the interactive theorem prover Isabelle/HOL. + +In particular, our semantics supports the following features of Solidity: +\begin{itemize} + \item \emph{Fixed-size integer types} of various lengths and corresponding arithmetic. + \item \emph{Domain-specific primitives}, such as money transfer or balance queries. + \item \emph{Different types of stores}, such as storage, memory, and stack. + \item \emph{Complex data types}, such as hash-maps and arrays. + \item \emph{Assignments with different semantics}, depending on data types. + \item An extendable \emph{gas model}. + \item \emph{Internal and external method calls}. +\end{itemize} + +A more abstract description of the semantics is given +in~\cite{marmsoler.ea:solidity-semantics:2021} and the conformance testing +approach for ensuring that our semantics conforms to the actual +implementation is described in~\cite{marmsoler.ea:conformance:2022}. + + +The rest of this document is automatically generated from the formalization in +Isabelle/HOL, i.e., all content is checked by Isabelle. The structure follows +the theory dependencies (see \autoref{fig:session-graph}). + +\begin{figure} + \centering + \includegraphics[height=.9\textheight, width=\textwidth,keepaspectratio]{session_graph} + \caption{The Dependency Graph of the Isabelle Theories.\label{fig:session-graph}} +\end{figure} + +\clearpage +%\input{session} + +\chapter{Preliminaries} +In this chapter, we discuss auxiliary formalizations and functions that +are used in our Solidity semantics but are more generic, i.e., not +specific to Solidity. This includes, for example, functions to convert +values of basic types to/from strings. + +\input{ReadShow.tex} +\input{Utils} +\input{StateMonad.tex} + +\chapter{Types and Accounts} +In this chapter, we discuss the basic data types of Solidity and the +representations of accounts. + +\input{Valuetypes.tex} +\input{Accounts} + +\chapter{Stores and Environment} +In this chapter, we focus on a particular aspect of Solidity that is +different to most programming languages: the handling of memory in +general and, in particular, the different between store and storage. + +\input{Storage.tex} +\input{Environment.tex} + +\chapter{Expressions and Statements} + +In this chapter, we formalize expressions, declarations, and +statements. The results up to here form the core of our Solidity +semantics. + +\input{Contracts.tex} +\input{Expressions.tex} +\input{Statements.tex} +\input{Solidity_Main.tex} + +\chapter{A Solidity Evaluation System} +This chapter discussed a tactic for symbolically executing Solidity statements +and expressions as well as provides a configuration for Isabelle's code +generator that allows us to generate an efficient implementation of our +executable formal semantics in, e.g., Haskell, SML, or Scala. In our test +framework, we use Haskell as a target language. +\input{Solidity_Symbex.tex} +\input{Solidity_Evaluator.tex} +\IfFileExists{Compile_Evaluator.tex}{% + \input{Compile_Evaluator.tex} +}{} + +\chapter{Verification Support} +This chapter presents a weakest precondition calculus and corresponding verification condition generator. + +\input{Weakest_Precondition.tex} + +\chapter{Applications} +In this chapter, we discuss various applications of our Solidity +semantics. + +\input{Reentrancy} +\input{Constant_Folding.tex} + +\IfFileExists{root.bib}{% + \bibliography{root} +}{} +\end{document} + +%%% Local Variables: +%%% mode: latex +%%% TeX-master: t +%%% End: