diff --git a/thys/Gale_Shapley/Gale_Shapley1.thy b/thys/Gale_Shapley/Gale_Shapley1.thy --- a/thys/Gale_Shapley/Gale_Shapley1.thy +++ b/thys/Gale_Shapley/Gale_Shapley1.thy @@ -1,2151 +1,2151 @@ (* Stepwise refinement of the Gale-Shapley algorithm down to executable functional code. Author: Tobias Nipkow *) section "Part 1: Refinement down to lists" theory Gale_Shapley1 imports Main "HOL-Hoare.Hoare_Logic" "List-Index.List_Index" "HOL-Library.While_Combinator" "HOL-Library.LaTeXsugar" begin subsection \Misc\ lemmas conj12 = conjunct1 conjunct2 syntax "_assign_list" :: "idt \ nat \ 'b \ 'com" ("(2_[_] :=/ _)" [70, 0, 65] 61) translations "xs[n] := e" \ "xs := CONST list_update xs n e" abbreviation upt_set :: "nat \ nat set" ("{<_}") where "{ {0.. 'a \ 'a \ bool" where "prefers P x y = (index P x < index P y)" abbreviation prefa :: "'a list \ 'a \ 'a \ bool" ("(_ \/ _ < _)" [50,50,50] 50) where "P \ x < y \ prefers P x y" lemma prefers_asym: "P \ x < y \ \ P \ y < x" by(simp add: prefers_def) lemma prefers_trans: "P \ x < y \ P \ y < z \ P \ x < z" by (meson order_less_trans prefers_def) fun rk_of_pref :: "nat \ nat list \ nat list \ nat list" where "rk_of_pref r rs (n#ns) = (rk_of_pref (r+1) rs ns)[n := r]" | "rk_of_pref r rs [] = rs" definition ranking :: "nat list \ nat list" where "ranking P = rk_of_pref 0 (replicate (length P) 0) P" lemma length_rk_of_pref[simp]: "length(rk_of_pref v vs P) = length vs" by(induction P arbitrary: v)(auto) lemma nth_rk_of_pref: "\ length P \ length rs; i \ set P; distinct P; set P \ { \ rk_of_pref r rs P ! i = index P i + r" by(induction P arbitrary: r i) (auto simp add: nth_list_update) lemma ranking_index: "\ length P = n; set P = { \ ranking P = map (index P) [0.. set P = { \ ranking P ! i < ranking P ! j \ P \ i < j" by(simp add: ranking_index prefers_def) subsection \Fixing the preference lists\ type_synonym prefs = "nat list list" locale Pref = fixes n fixes P :: prefs fixes Q :: prefs defines "n \ length P" assumes length_Q: "length Q = n" assumes P_set: "a < n \ length(P!a) = n \ set(P!a) = { length(Q!b) = n \ set(Q!b) = { bool" where "wf xs \ length xs = n \ set xs \ { wf A; a < n \ \ A!a < n" by (simp add: subset_eq) corollary wf_le_n1: "\ wf A; a < n \ \ A!a \ n-1" using wf_less_n by fastforce lemma sumA_ub: "wf A \ (\a n*(n-1)" using sum_bounded_above[of "{..The (termination) variant(s)\ text \Basic idea: either some \A!a\ is incremented or size of \M\ is incremented, but this cannot go on forever because in the worst case all \A!a = n-1\ and \M = n\. Because \n*(n-1) + n = n^2\, this leads to the following simple variant:\ definition var0 :: "nat list \ nat set \ nat" where [simp]: "var0 A M = (n^2 - ((\a { a \ M" shows "var0 A (M \ {a}) < var0 A M" proof - have 2: "M \ { { {a n*n" using sumA_ub[OF assms(1)] 0 by (simp add: algebra_simps le_diff_conv2) have "var0 (A[a' := A ! a' + 1]) M = n*n - (1 + (A ! a' + sum ((!) A) ({ = n^2 - (1 + (\a < n^2 - ((\a nat set \ nat" where [simp]: "var A M = (n^2 - n + 1 - ((\a n-1" "\a < n. a \ a' \ A!a \ n-2" shows "(\a (n-1)*(n-1)" proof - have "(\aa \ ({ {a'}. A!a)" by (simp add: assms(1) atLeast0LessThan insert_absorb) also have "\ =(\a \ { \ (\a \ { \ (n-1)*(n-2) + (n-1)" using sum_bounded_above[of "{.. = (n-1)*(n-1)" by (metis Suc_diff_Suc Suc_eq_plus1 add.commute diff_is_0_eq' linorder_not_le mult_Suc_right mult_cancel_left nat_1_add_1) finally show ?thesis . qed definition "match A a = P ! a ! (A ! a)" lemma match_less_n: "\ wf A; a < n \ \ match A a < n" by (metis P_set atLeastLessThan_iff match_def nth_mem subset_eq) lemma match_upd_neq: "\ wf A; a < n; a \ a' \ \ match (A[a := b]) a' = match A a'" by (simp add: match_def) definition stable :: "nat list \ nat set \ bool" where "stable A M = (\(\a\M. \a'\M. P ! a \ match A a' < match A a \ Q ! match A a' \ a < a'))" text \The set of Bs that an A would prefer to its current match, i.e. all Bs above its current match \A!a\.\ abbreviation preferred where "preferred A a == nth (P!a) ` { inj_on (match A) M)" text \If \a'\ is unmatched and final then all other \a\ are matched:\ lemma final_last: assumes M: "M \ { match A ` M" and a: "a < n \ a \ M" and final: "A ! a + 1 = n" shows "insert a M = { card M" by(rule surj_card_le[OF subset_eq_atLeast0_lessThan_finite[OF M] pref_match']) have 3: "card M < n" using M a by (metis atLeast0LessThan card_seteq order.refl finite_atLeastLessThan le_neq_implies_less lessThan_iff subset_eq_atLeast0_lessThan_card) have "Suc (card M) = n" using 1 2 3 by simp thus ?thesis using M a by (simp add: card_subset_eq finite_subset) qed lemma more_choices: assumes A: "wf A" and M: "M \ { { match A ` M" and "a < n" and matched: "match A a \ match A ` M" shows "A ! a + 1 < n" proof (rule ccontr) assume "\ A ! a + 1 < n" hence "A ! a + 1 = n" using A \a < n\ unfolding matching_def by (metis add.commute wf_less_n linorder_neqE_nat not_less_eq plus_1_eq_Suc) hence *: "nth (P ! a) ` { match A ` M" using pref_match' matched less_Suc_eq match_def by fastforce have "nth (P!a) ` {a < n\] by (metis map_nth set_map set_upt) hence "{ match A ` M" using * by metis hence "card { card M" using finite_subset[OF \M \ { finite_atLeastLessThan] by (metis surj_card_le) then show False using M card_seteq by blast qed corollary more_choices_matched: assumes "wf A" "M \ { { match A ` M" and "a \ M" shows "A ! a + 1 < n" using more_choices[OF assms(1-4)] \a \ M\ \M \ { atLeastLessThan_iff by blast lemma atmost1_final: assumes M: "M \ {a match A ` M" shows "\\<^sub>\\<^sub>1 a. a < n \ a \ M \ A ! a + 1 = n" apply rule subgoal for x y using final_last[OF M inj, of x] final_last[OF M inj, of y] assms(3) by blast done lemma sumA_UB: assumes "matching A M" "M \ { {a match A ` M" shows "(\a (n-1)^2" proof - have "wf A" using assms(1) by(simp) have M: "\a\M. A!a + 1 < n" using more_choices_matched[OF \wf A\ assms(2-3)] assms(4) \M \ { atLeastLessThan_iff by blast note Ainj = conj12[OF assms(1)[unfolded matching_def]] show ?thesis proof (cases "\a' M \ A!a' + 1 = n") case True then obtain a' where a': "a' M" "A!a' + 1 = n" using \M \ { \M \ { by blast hence "\a a' \ A!a \ n-2" using Uniq_D[OF atmost1_final[OF assms(2) Ainj(2) assms(4)], of a'] M wf_le_n1[OF Ainj(1)] by (metis Suc_1 Suc_eq_plus1 add_diff_cancel_right' add_le_imp_le_diff diff_diff_left less_eq_Suc_le order_less_le) from sumA_ub2[OF a'(1) _ this] a'(3) show ?thesis unfolding power2_eq_square by linarith next case False hence "\a' M \ A ! a' + 1 < n" by (metis Suc_eq_plus1 Suc_lessI wf_less_n[OF Ainj(1)]) with M have "\aa n*(n-2)" using sum_bounded_above[of "{.. \ (n-1)*(n-1)" by(simp add: algebra_simps) finally show ?thesis unfolding power2_eq_square . qed qed lemma var_ub: assumes "matching A M" "M \ { {a match A ` M" shows "(\a { n^2 + 1 - 2*n" using sumA_UB[OF assms(1-4)] by (simp add: power2_eq_square algebra_simps) have 4: "2*n \ Suc (n^2)" using le_square[of n] unfolding power2_eq_square by (metis Suc_1 add_mono_thms_linordered_semiring(1) le_SucI mult_2 mult_le_mono1 not_less_eq_eq plus_1_eq_Suc) show "(\a { {a match A ` M" "a \ M" shows "var A (M \ {a}) < var A M" proof - have 2: "M \ { n^2 + 1 - 2*n" using sumA_UB[OF assms(1-4)] by (simp add: power2_eq_square algebra_simps) have 5: "2*n \ Suc (n^2)" using le_square[of n] unfolding power2_eq_square by (metis Suc_1 add_mono_thms_linordered_semiring(1) le_SucI mult_2 mult_le_mono1 not_less_eq_eq plus_1_eq_Suc) have 6: "(\aa \ M\ finite_subset[OF assms(2)] by(simp) qed lemma var_next: assumes"matching A M" "M \ { {a match A ` M" "a < n" shows "var (A[a := A ! a + 1]) M < var A M" proof - have "var (A[a := A ! a + 1]) M = n*n - n + 1 - (1 + (A ! a + sum ((!) A) ({ = n^2 - n + 1 - (1 + (\a < n^2 - n + 1 - ((\aAuxiliary Predicates\ text \The following two predicates express the same property: if \a\ prefers \b\ over \a\'s current match, then \b\ is matched with an \a'\ that \b\ prefers to \a\.\ definition pref_match where "pref_match A M = (\ab b < match A a \ (\a'\M. b = match A a' \ Q ! b \ a' < a))" definition pref_match' where "pref_match' A M = (\ab \ preferred A a. \a'\M. b = match A a' \ Q ! b \ a' < a)" lemma pref_match'_iff: "wf A \ pref_match' A M = pref_match A M" apply (auto simp add: pref_match'_def pref_match_def imp_ex prefers_def match_def) apply (smt (verit) P_set atLeast0LessThan order.strict_trans index_first lessThan_iff linorder_neqE_nat nth_index) by (smt (verit, best) P_set atLeast0LessThan card_atLeastLessThan card_distinct diff_zero in_mono index_nth_id lessThan_iff less_trans nth_mem) definition opti\<^sub>a where "opti\<^sub>a A = (\A'. matching A' { stable A' { (\a match A' a < match A a))" definition pessi\<^sub>b where "pessi\<^sub>b A = (\A'. matching A' { stable A' { (\aa' Q ! match A a \ a < a'))" lemma opti\<^sub>a_pessi\<^sub>b: assumes "opti\<^sub>a A" shows "pessi\<^sub>b A" unfolding pessi\<^sub>b_def proof (safe, goal_cases) case (1 A' a a') have "\ P!a \ match A a < match A' a" using 1 by (metis atLeast0LessThan lessThan_iff stable_def) with 1 \opti\<^sub>a A\ show ?case using P_set match_less_n opti\<^sub>a_def prefers_def unfolding matching_def by (metis (no_types) atLeast0LessThan inj_on_contraD lessThan_iff less_not_refl linorder_neqE_nat nth_index) qed lemma opti\<^sub>a_inv: assumes A: "wf A" and a: "a < n" and a': "a' < n" and same_match: "match A a' = match A a" and pref: "Q ! match A a' \ a' < a" and "opti\<^sub>a A" shows "opti\<^sub>a (A[a := A ! a + 1])" proof (unfold opti\<^sub>a_def matching_def, rule notI, elim exE conjE) note opti\<^sub>a = \opti\<^sub>a A\[unfolded opti\<^sub>a_def matching_def] let ?A = "A[a := A ! a + 1]" fix A' a'' assume "a'' < n" and A': "length A' = n" "set A' \ { match A' a'' < match ?A a''" show False proof cases assume [simp]: "a'' = a" have "A!a < n" using A a by(simp add: subset_eq) with A A' a pref_a'' have "P ! a \ match A' a < match A a \ match A' a = match A a" apply(auto simp: prefers_def match_def) by (smt (verit) P_set wf_less_n card_atLeastLessThan card_distinct diff_zero index_nth_id not_less_eq not_less_less_Suc_eq) thus False proof assume "P ! a \ match A' a < match A a " thus False using opti\<^sub>a A' \a < n\ by(fastforce) next assume "match A' a = match A a" have "a \ a'" using pref a' by(auto simp: prefers_def) hence "P ! a' \ match A' a < match A' a' \ Q ! match A' a \ a' < a" using opti\<^sub>a pref A' same_match \match A' a = match A a\ a a' by (metis P_set atLeast0LessThan match_less_n inj_onD lessThan_iff linorder_neqE_nat nth_index prefers_def) thus False using a a' \a \ a'\ A'(3) by (metis stable_def atLeastLessThan_iff zero_le) qed next assume "a'' \ a" thus False using opti\<^sub>a A' pref_a'' \a'' < n\ by(metis match_def nth_list_update_neq) qed qed lemma pref_match_stable: "\ matching A { \ stable A {Algorithm 1\ definition invAM where [simp]: "invAM A M = (matching A M \ M \ { pref_match A M \ opti\<^sub>a A)" lemma invAM_match: "\ invAM A M; a < n \ a \ M; match A a \ match A ` M \ \ invAM A (M \ {a})" by(simp add: pref_match_def) lemma invAM_swap: assumes "invAM A M" assumes a: "a < n \ a \ M" and a': "a' \ M \ match A a' = match A a" and pref: "Q ! match A a' \ a < a'" shows "invAM (A[a' := A!a'+1]) (M - {a'} \ {a})" proof - have A: "wf A" and M : "M \ {a A" by(insert \invAM A M\) (auto) have "M \ { a'" using a' a M by auto have pref_match': "pref_match' A M" using pref_match pref_match'_iff[OF A] by blast let ?A = "A[a' := A!a'+1]" let ?M = "M - {a'} \ {a}" have neq_a': "\x. x \ ?M \ a' \ x" using \a \ a'\ by blast have \set ?A \ { apply(rule set_update_subsetI[OF A[THEN conjunct2]]) using more_choices[OF _ M \M \ {] A inj pref_match' a' subsetD[OF M, of a'] by(fastforce simp: pref_match'_def) hence "wf ?A" using A by(simp) moreover have "inj_on (match ?A) ?M" using a a' inj by(simp add: match_def inj_on_def)(metis Diff_iff insert_iff nth_list_update_neq) moreover have "pref_match' ?A ?M" using a a' pref_match' A pref \a' < n\ apply(simp add: pref_match'_def match_upd_neq neq_a' Ball_def Bex_def image_iff imp_ex nth_list_update less_Suc_eq flip: match_def) by (metis prefers_trans) moreover have "opti\<^sub>a ?A" using opti\<^sub>a_inv[OF A \a' < n\ _ _ _ \opti\<^sub>a A\] a a'[THEN conjunct2] pref by auto ultimately show ?thesis using a a' M pref_match'_iff by auto qed (* TODO: could also be used in invAM_next *) lemma preferred_subset_match_if_invAM: assumes "invAM A M" shows "\a match A ` M" (is ?P) proof - have A: "wf A" and pref_match: "pref_match A M" using assms(1) by auto note pref_match' = pref_match[THEN pref_match'_iff[OF A, THEN iffD2]] thus pref_match1: "\a match A ` M" unfolding pref_match'_def by blast qed lemma invAM_next: assumes "invAM A M" assumes a: "a < n \ a \ M" and a': "a' \ M \ match A a' = match A a" and pref: "\ Q ! match A a' \ a < a'" shows "invAM (A[a := A!a + 1]) M" proof - have A: "wf A" and M : "M \ {a: "opti\<^sub>a A" and "a' < n" by(insert \invAM A M\ a') (auto) hence pref': "Q ! match A a' \ a' < a" using pref a a' Q_set unfolding prefers_def by (metis match_def match_less_n index_eq_index_conv linorder_less_linear subsetD) have "M \ {x. x\ M \ a \ x" using a by blast have pref_match': "pref_match' A M" using pref_match pref_match'_iff[OF A,of M] by blast hence "\a match A ` M" unfolding pref_match'_def by blast hence "A!a + 1 < n" using more_choices[OF _ M \M \ {] A inj a a' unfolding matching_def by (metis (no_types, lifting) imageI) let ?A = "A[a := A!a+1]" have "wf ?A" using A \A!a + 1 < n\ by(simp add: set_update_subsetI) moreover have "inj_on (match ?A) M" using a inj by(simp add: match_def inj_on_def) (metis nth_list_update_neq) moreover have "pref_match' ?A M" using a pref_match' pref' A a' neq_a by(auto simp: match_upd_neq pref_match'_def Ball_def Bex_def image_iff nth_list_update imp_ex less_Suc_eq simp flip: match_def) moreover have "opti\<^sub>a ?A" using opti\<^sub>a_inv[OF A conjunct1[OF a] \a' < n\ conjunct2[OF a'] pref' opti\<^sub>a] . ultimately show ?thesis using M by (simp add: pref_match'_iff) qed lemma Gale_Shapley1: "VARS M A a a' b [M = {} \ A = replicate n 0] WHILE M \ { a \ M); b := match A a; IF b \ match A ` M THEN M := M \ {a} ELSE a' := (SOME a'. a' \ M \ match A a' = b); IF Q ! match A a' \ a < a' THEN A[a'] := A!a'+1; M := M - {a'} \ {a} ELSE A[a] := A!a+1 FI FI OD [matching A { stable A { opti\<^sub>a A]" proof (vcg_tc, goal_cases) case 1 thus ?case by(auto simp: stable_def opti\<^sub>a_def pref_match_def P_set card_distinct match_def index_nth_id prefers_def) next case 3 thus ?case using pref_match_stable by auto next case (2 v M A a) hence invAM: "invAM A M" and m: "matching A M" and M: "M \ { {a A" and v: "var A M = v" by auto note Ainj = conj12[OF m[unfolded matching_def]] note pref_match1 = preferred_subset_match_if_invAM[OF invAM] define a where "a = (SOME a. a < n \ a \ M)" have a: "a < n \ a \ M" unfolding a_def using M by (metis (no_types, lifting) atLeastLessThan_iff someI_ex subsetI subset_antisym) show ?case (is "?P((SOME a. a < n \ a \ M))") unfolding a_def[symmetric] proof - show "?P a" (is "(?not_matched \ ?THEN) \ (\ ?not_matched \ ?ELSE)") proof (rule; rule) assume ?not_matched show ?THEN proof(simp only:mem_Collect_eq prod.case, rule conjI, goal_cases) case 1 show ?case using invAM_match[OF invAM a \?not_matched\] . case 2 show ?case using var_match[OF m M pref_match1] var0_match[OF Ainj(1) M(1)] a unfolding v by blast qed next assume matched: "\ ?not_matched" define a' where "a' = (SOME a'. a' \ M \ match A a' = match A a)" have a': "a' \ M \ match A a' = match A a" unfolding a'_def using matched by (smt (verit) image_iff someI_ex) hence "a' < n" "a \ a'" using a M atLeast0LessThan by auto show ?ELSE (is "?P((SOME a'. a' \ M \ match A a' = match A a))") unfolding a'_def[symmetric] proof - show "?P a'" (is "(?pref \ ?THEN) \ (\ ?pref \ ?ELSE)") proof (rule; rule) assume ?pref show ?THEN proof (simp only: mem_Collect_eq prod.case, rule conjI, goal_cases) case 1 show ?case by(rule invAM_swap[OF invAM a a' \?pref\]) case 2 have "card(M - {a'} \ {a}) = card M" using a a' card.remove subset_eq_atLeast0_lessThan_finite[OF M(1)] by fastforce thus ?case using v var_next[OF m M pref_match1 \a' < n\] var0_next[OF Ainj(1) M \a' < n\] by simp qed next assume "\ ?pref" show ?ELSE proof (simp only: mem_Collect_eq prod.case, rule conjI, goal_cases) case 1 show ?case using invAM_next[OF invAM a a' \\ ?pref\] . case 2 show ?case using a v var_next[OF m M pref_match1, of a] var0_next[OF Ainj(1) M, of a] by simp qed qed qed qed qed qed text \Proof also works for @{const var0} instead of @{const var}.\ subsection \Algorithm 2: List of unmatched As\ abbreviation invas where "invas as == (set as \ { distinct as)" lemma Gale_Shapley2: "VARS A a a' as b [as = [0.. A = replicate n 0] WHILE as \ [] INV { invAM A ({ invas as} VAR {var A ({ match A ` ({ { match A a' = b); IF Q ! match A a' \ a < a' THEN A[a'] := A!a'+1; as := a' # tl as ELSE A[a] := A!a+1 FI FI OD [matching A { stable A { opti\<^sub>a A]" proof (vcg_tc, goal_cases) case 1 thus ?case by(auto simp: stable_def opti\<^sub>a_def pref_match_def P_set card_distinct match_def index_nth_id prefers_def) next case 3 thus ?case using pref_match_stable by auto next case (2 v A _ a' as) let ?M = "{ { []" and as: "invas as" and v: "var A ?M = v" using 2 by auto note pref_match1 = preferred_subset_match_if_invAM[OF invAM] from \as \ []\ obtain a as' where aseq: "as = a # as'" by (fastforce simp: neq_Nil_conv) have set_as: "?M \ {a} = { a \ ?M" using as unfolding aseq by (simp) show ?case proof (simp only: aseq list.sel, goal_cases) case 1 show ?case (is "(?not_matched \ ?THEN) \ (\ ?not_matched \ ?ELSE)") proof (rule; rule) assume ?not_matched then have nm: "match A a \ match A ` ?M" unfolding aseq . show ?THEN proof(simp only:mem_Collect_eq prod.case, rule conjI, goal_cases) case 1 show ?case using invAM_match[OF invAM a nm] as unfolding set_as by (simp add: aseq) case 2 show ?case using var_match[OF m M _ pref_match1, of a] a atLeast0LessThan unfolding set_as v by blast qed next assume matched: "\ ?not_matched" define a' where "a' = (SOME a'. a' \ ?M \ match A a' = match A a)" have a': "a' \ ?M \ match A a' = match A a" unfolding a'_def aseq using matched by (smt (verit) image_iff someI_ex) hence "a' < n" "a \ a'" using a M atLeast0LessThan by auto show ?ELSE unfolding aseq[symmetric] a'_def[symmetric] proof (goal_cases) case 1 show ?case (is "(?pref \ ?THEN) \ (\ ?pref \ ?ELSE)") proof (rule; rule) assume ?pref show ?THEN proof (simp only: mem_Collect_eq prod.case, rule conjI, goal_cases) have *: "{ {a} = {?pref\] unfolding * using a' as aseq by force case 2 have "card({a' < n\ a atLeast0LessThan by (metis Suc_eq_plus1 lessThan_iff var_def) qed next assume "\ ?pref" show ?ELSE proof (simp only: mem_Collect_eq prod.case, rule conjI, goal_cases) case 1 show ?case using invAM_next[OF \invAM A ?M\ a a' \\ ?pref\] as by blast case 2 show ?case using a v var_next[OF m M _ pref_match1, of a] by (metis Suc_eq_plus1 atLeast0LessThan lessThan_iff) qed qed qed qed qed qed subsection \Algorithm 3: Record matching of Bs to As\ abbreviation invAB :: "nat list \ (nat \ nat option) \ nat set \ bool" where "invAB A B M == (ran B = M \ (\b a. B b = Some a \ match A a = b))" lemma invAB_swap: assumes invAB: "invAB A B M" assumes a: "a < n \ a \ M" and a': "a' \ M \ match A a' = match A a" and "inj_on B (dom B)" "B(match A a) = Some a'" shows "invAB (A[a' := A!a'+1]) (B(match A a := Some a)) (M - {a'} \ {a})" proof - have "\b x. b \ match A a \ B b = Some x \ a'\ x" using invAB a' by blast moreover have "a \ a'" using a a' by auto ultimately show ?thesis using assms by(simp add: ran_map_upd_Some match_def) qed lemma Gale_Shapley3: "VARS A B a a' as b [as = [0.. A = replicate n 0 \ B = (\_. None)] WHILE as \ [] INV { invAM A ({ invAB A B ({ invas as} VAR {var A ({ a < a' THEN B := B(b := Some a); A[a'] := A!a'+1; as := a' # tl as ELSE A[a] := A!a+1 FI FI OD [matching A { stable A { opti\<^sub>a A]" proof (vcg_tc, goal_cases) case 1 thus ?case by(auto simp: stable_def opti\<^sub>a_def pref_match_def P_set card_distinct match_def index_nth_id prefers_def) next case 3 thus ?case using pref_match_stable by auto next case (2 v A B _ a' as) let ?M = "{ { []" and as: "invas as" and invAB: "invAB A B ?M" and v: "var A ?M = v" using 2 by auto note pref_match1 = preferred_subset_match_if_invAM[OF invAM] from \as \ []\ obtain a as' where aseq: "as = a # as'" by (fastforce simp: neq_Nil_conv) have set_as: "?M \ {a} = { a \ ?M" using as unfolding aseq by (simp) show ?case proof (simp only: aseq list.sel, goal_cases) case 1 show ?case (is "(?not_matched \ ?THEN) \ (\ ?not_matched \ ?ELSE)") proof (rule; rule) assume ?not_matched then have nm: "match A a \ match A ` ?M" using invAB unfolding aseq ran_def apply (clarsimp simp: set_eq_iff) using not_None_eq by blast show ?THEN proof(simp only:mem_Collect_eq prod.case, rule conjI, goal_cases) have invAM': "invAM A ({?not_matched\ set_as by (simp) case 1 show ?case using invAM' as invAB' unfolding set_as aseq by (metis distinct.simps(2) insert_subset list.simps(15)) case 2 show ?case using var_match[OF m M _ pref_match1, of a] a atLeast0LessThan unfolding set_as v by blast qed next assume matched: "\ ?not_matched" then obtain a' where a'eq: "B(match A a) = Some a'" by auto have a': "a' \ ?M \ match A a' = match A a" unfolding aseq using a'eq invAB by (metis ranI aseq) hence "a' < n" "a \ a'" using a M atLeast0LessThan by auto show ?ELSE unfolding aseq[symmetric] a'eq option.sel proof (goal_cases) have inj_dom: "inj_on B (dom B)" by (metis (mono_tags) domD inj_onI invAB) case 1 show ?case (is "(?pref \ ?THEN) \ (\ ?pref \ ?ELSE)") proof (rule; rule) assume ?pref show ?THEN proof (simp only: mem_Collect_eq prod.case, rule conjI, goal_cases) have *: "{ {a} = {b x. b \ match A a \ B b = Some x \ a'\ x" using invAB a' by blast have invAB': "invAB (A[a' := A ! a' + 1]) (B(match A a := Some a)) ({?pref\] invAB' unfolding * using a' as aseq by simp case 2 have "card({a' < n\ a atLeast0LessThan by (metis Suc_eq_plus1 lessThan_iff var_def) qed next assume "\ ?pref" show ?ELSE proof (simp only: mem_Collect_eq prod.case, rule conjI, goal_cases) case 1 have "invAB (A[a := A ! a + 1]) B ?M" using invAB a by (metis match_def nth_list_update_neq ranI) thus ?case using invAM_next[OF invAM a a' \\ ?pref\] as by blast case 2 show ?case using a v var_next[OF m M _ pref_match1, of a] by (metis Suc_eq_plus1 atLeast0LessThan lessThan_iff) qed qed qed qed qed qed subsection \Unused data refinement step\ (* begin unused: directly implement function B via lists B and M (data refinement); also done in Alg. 4 in a more principled manner *) abbreviation invAB' :: "nat list \ nat list \ bool list \ nat set \ bool" where "invAB' A B M M' == (length B = n \ length M = n \ M' = nth B ` {b. b < n \ M!b} \ (\b B!b < n \ match A (B!b) = b))" lemma Gale_Shapley4_unused: "VARS A B M a a' as b [as = [0.. A = replicate n 0 \ B = replicate n 0 \ M = replicate n False] WHILE as \ [] INV { invAM A ({ invAB' A B M ({ invas as} VAR {var A ({ (M ! b) THEN M[b] := True; B[b] := a; as := tl as ELSE a' := B ! b; IF Q ! match A a' \ a < a' THEN B[b] := a; A[a'] := A!a'+1; as := a' # tl as ELSE A[a] := A!a+1 FI FI OD [wf A \ inj_on (match A) { stable A { opti\<^sub>a A]" proof (vcg_tc, goal_cases) case 1 thus ?case by(auto simp: stable_def opti\<^sub>a_def pref_match_def P_set card_distinct match_def index_nth_id prefers_def) next case 3 thus ?case using pref_match_stable by auto next case (2 v A B M _ a' as) let ?M = "{ { []" and as: "invas as" and invAB: "invAB' A B M ?M" and v: "var A ?M = v" using 2 by auto note pref_match1 = preferred_subset_match_if_invAM[OF invAM] from notall obtain a as' where aseq: "as = a # as'" by (fastforce simp: neq_Nil_conv) have set_as: "?M \ {a} = { a \ ?M" using as unfolding aseq by (simp) show ?case proof (simp only: aseq list.sel, goal_cases) case 1 show ?case (is "(?not_matched \ ?THEN) \ (\ ?not_matched \ ?ELSE)") proof (rule; rule) assume ?not_matched then have nm: "match A a \ match A ` ?M" using invAB set_as unfolding aseq by force show ?THEN proof(simp only:mem_Collect_eq prod.case, rule conjI, goal_cases) have invAM': "invAM A ({?not_matched\ set_as match_less_n[OF A] a by (auto simp add: image_def nth_list_update) case 1 show ?case using invAM' invAB as invAB' unfolding set_as aseq by (metis distinct.simps(2) insert_subset list.simps(15)) case 2 show ?case using var_match[OF m M _ pref_match1, of a] a atLeast0LessThan unfolding set_as v by blast qed next assume matched: "\ ?not_matched" hence "match A a \ match A ` ({ ?M \ match A ?a = match A a" using invAB match_less_n[OF A] matched a by blast hence "?a < n" "a \ ?a" using a by auto show ?ELSE unfolding aseq option.sel proof (goal_cases) case 1 show ?case (is "(?pref \ ?THEN) \ (\ ?pref \ ?ELSE)") proof (rule; rule) assume ?pref show ?THEN proof (simp only: mem_Collect_eq prod.case, rule conjI, goal_cases) have *: "{ {a} = {b match A a \ M!b \ ?a \ B!b" using invAB a' by metis have invAB': "invAB' (A[?a := A ! ?a + 1]) (B[match A a := a]) M ({a \ ?a\ a' match_less_n[OF A, of a] a apply (simp add: nth_list_update) apply rule apply(auto simp add: image_def)[] apply (clarsimp simp add: match_def) apply (metis (opaque_lifting) nth_list_update_neq) done case 1 show ?case using invAM_swap[OF invAM a a' \?pref\] invAB' unfolding * using a' as aseq by (auto) case 2 have "card({?a < n\ a atLeast0LessThan by (metis Suc_eq_plus1 lessThan_iff var_def) qed next assume "\ ?pref" show ?ELSE proof (simp only: mem_Collect_eq prod.case, rule conjI, goal_cases) case 1 have "invAB' (A[a := A ! a + 1]) B M ({a \ ?a\ by (metis match_def nth_list_update_neq) thus ?case using invAM_next[OF invAM a a' \\ ?pref\] as aseq by fastforce case 2 show ?case using a v var_next[OF m M _ pref_match1, of a] aseq by (metis Suc_eq_plus1 atLeast0LessThan lessThan_iff) qed qed qed qed qed qed subsection \Algorithm 4: remove list of unmatched As\ subsubsection \An initial version\ text \The inner variant appears intuitive but complicates the derivation of an overall complexity bound because the inner variant also depends on a variable that is modified by the outer loop.\ lemma Gale_Shapley4: "VARS A B ai a a' [ai = 0 \ A = replicate n 0 \ B = (\_. None)] WHILE ai < n INV { invAM A { invAB A B { ai \ n } VAR {z = n - ai} DO a := ai; WHILE B (match A a) \ None INV { invAM A ({ invAB A B ({ (a \ ai \ ai < n) \ z = n-ai } VAR {var A { a < a' THEN B := B(match A a := Some a); A[a'] := A!a'+1; a := a' ELSE A[a] := A!a+1 FI OD; B := B(match A a := Some a); ai := ai+1 OD [matching A { stable A { opti\<^sub>a A]" proof (vcg_tc, goal_cases) case 1 thus ?case (* outer invar holds initially *) by(auto simp: stable_def pref_match_def P_set card_distinct match_def index_nth_id prefers_def opti\<^sub>a_def)[] next case 2 (* outer invar and b implies inner invar *) thus ?case by (auto simp: atLeastLessThanSuc_atLeastAtMost simp flip: atLeastLessThan_eq_atLeastAtMost_diff) next case (4 z A B ai a) (* inner invar' and not b implies outer invar *) note inv = 4[THEN conjunct1] note invAM = inv[THEN conjunct1] note aai = inv[THEN conjunct2,THEN conjunct2] show ?case proof (simp only: mem_Collect_eq prod.case, rule conjI, goal_cases) case 1 have *: "{ match A a \ match A ` ({ match A ` ({ { None" and as: "a \ ai \ ai < n" and invAB: "invAB A B ?M" and v: "var A ?M = v" using 3 by auto note invar = 3[THEN conjunct1,THEN conjunct1] note pref_match1 = preferred_subset_match_if_invAM[OF invAM] from matched obtain a' where a'eq: "B(match A a) = Some a'" by auto have a': "a' \ ?M \ match A a' = match A a" using a'eq invAB by (metis ranI) have a: "a < n \ a \ ?M" using invar by auto have "?M \ { ?THEN) \ (\ ?unstab \ ?ELSE)") proof (rule; rule) assume ?unstab have *: "{ {a} = { a)) ({?unstab\] invAB' invar a' unfolding * by (simp) next case 2 show ?case using v var_next[OF m M \?M \ { pref_match1 \a' < n\] card by (metis var_def Suc_eq_plus1 psubset_eq) qed next assume "\ ?unstab" show ?ELSE proof (simp only: mem_Collect_eq prod.case, rule conjI, goal_cases) have *: "\b a'. B b = Some a' \ a \ a'" by (metis invAB ranI a) case 1 show ?case using invAM_next[OF invAM a a' \\ ?unstab\] invar * by (simp add: match_def) next case 2 show ?case using v var_next[OF m M \?M \ { pref_match1, of a] a card by (metis Suc_eq_plus1 var_def) qed qed qed qed subsubsection \A better inner variant\ text \This is the definitive version of Algorithm 4. The inner variant is changed to support the easy derivation of the precise upper bound of the number of executed actions. This variant shows that the algorithm can at most execute @{term "n^2 - n + 1"} basic actions (match, swap, next).\ definition var2 :: "nat list \ nat" where [simp]: "var2 A = (n-1)^2 - (\aBecause \A\ is not changed by the outer loop, the initial value of @{term "var2 A"}, which is @{term "(n-1)^2"}, is an upper bound of the number of iterations of the inner loop. To this we need to add \n\ because the outer loop executes additional \n\ match actions at the end of the loop body. Thus at most @{prop "(n-1)^2 + n = n^2 - n + 1"} actions are executed, exactly as in the earlier algorithms.\ lemma var2_next: assumes "invAM (A[a := A!a + 1]) M" "M \ {aaawf ?A\ \a < n\ by(simp add: nth_list_update sum.If_cases lessThan_atLeast0 flip:Diff_eq) also have "\ = (\aa < n\ member_le_sum[of a "{a (n-1)^2" using sumA_UB[of ?A M] assms(1,2) by (meson invAM_def preferred_subset_match_if_invAM) with 1 show ?thesis unfolding var2_def by auto qed lemma Gale_Shapley4_var2: "VARS A B ai a a' [ai = 0 \ A = replicate n 0 \ B = (\_. None)] WHILE ai < n INV { invAM A { invAB A B { ai \ n } VAR {z = n - ai} DO a := ai; WHILE B (match A a) \ None INV { invAM A ({ invAB A B ({ (a \ ai \ ai < n) \ z = n-ai } VAR {var2 A} DO a' := the(B (match A a)); IF Q ! match A a' \ a < a' THEN B := B(match A a := Some a); A[a'] := A!a'+1; a := a' ELSE A[a] := A!a+1 FI OD; B := B(match A a := Some a); ai := ai+1 OD [matching A { stable A { opti\<^sub>a A]" proof (vcg_tc, goal_cases) case 1 thus ?case (* outer invar holds initially *) by(auto simp: stable_def pref_match_def P_set card_distinct match_def index_nth_id prefers_def opti\<^sub>a_def)[] next case 2 (* outer invar and b implies inner invar *) thus ?case by (auto simp: atLeastLessThanSuc_atLeastAtMost simp flip: atLeastLessThan_eq_atLeastAtMost_diff) next case (4 z A B ai a) (* inner invar' and not b implies outer invar *) note inv = 4[THEN conjunct1] note invAM = inv[THEN conjunct1] note aai = inv[THEN conjunct2,THEN conjunct2] show ?case proof (simp only: mem_Collect_eq prod.case, rule conjI, goal_cases) case 1 have *: "{ match A a \ match A ` ({ match A ` ({ { None" and as: "a \ ai \ ai < n" and invAB: "invAB A B ?M" and v: "var2 A = v" using 3 by auto note invar = 3[THEN conjunct1,THEN conjunct1] from matched obtain a' where a'eq: "B(match A a) = Some a'" by auto have a': "a' \ ?M \ match A a' = match A a" using a'eq invAB by (metis ranI) have a: "a < n \ a \ ?M" using invar by auto have "?M \ { ?THEN) \ (\ ?unstab \ ?ELSE)") proof (rule; rule) assume ?unstab have *: "{ {a} = {?unstab\] have inj_dom: "inj_on B (dom B)" by (metis (mono_tags) domD inj_onI invAB) have invAB': "invAB (A[a' := A ! a' + 1]) (B(match A a \ a)) ({a' < n\ * atLeast0LessThan by auto qed next assume "\ ?unstab" show ?ELSE proof (simp only: mem_Collect_eq prod.case, rule conjI, goal_cases) note invAM' = invAM_next[OF invAM a a' \\ ?unstab\] have *: "\b a'. B b = Some a' \ a \ a'" by (metis invAB ranI a) case 1 show ?case using invAM' invar * by (simp add: match_def) case 2 show ?case using v var2_next[OF invAM'] a \?M \ { by blast qed qed qed qed subsubsection \Algorithm 4.1: single-loop variant\ text \A bit of a relic because it is an instance of a general program transformation for merging nested loops by an additional test inside the single loop.\ lemma Gale_Shapley4_1: "VARS A B a a' ai b [ai = 0 \ a = 0 \ A = replicate n 0 \ B = (\_. None)] WHILE ai < n INV { invAM A ({ invAB A B ({ (a \ ai \ ai \ n) \ (ai=n \ a=ai)} VAR {var A ({ a < a' THEN B := B(b := Some a); A[a'] := A!a'+1; a := a' ELSE A[a] := A!a+1 FI FI OD [matching A { stable A { opti\<^sub>a A]" proof (vcg_tc, goal_cases) case 1 thus ?case by(auto simp: stable_def opti\<^sub>a_def pref_match_def P_set card_distinct match_def index_nth_id prefers_def) next case 3 thus ?case using pref_match_stable using atLeast0_lessThan_Suc by force next case (2 v A B a a' ai b) let ?M = "{ { ai \ ai < n" and invAB: "invAB A B ?M" and v: "var A ?M = v" using 2 by auto note invar = 2[THEN conjunct1,THEN conjunct1] note pref_match1 = preferred_subset_match_if_invAM[OF invAM] have a: "a < n \ a \ ?M" using as by (simp) show ?case (is "(?not_matched \ ?THEN) \ (\ ?not_matched \ ?ELSE)") proof (rule; rule) assume ?not_matched then have nm: "match A a \ match A ` ?M" using invAB unfolding ran_def apply (clarsimp simp: set_eq_iff) using not_None_eq by blast show ?THEN proof(simp only:mem_Collect_eq prod.case, rule conjI, goal_cases) have *: "{ {a} = {?not_matched\ ** by (simp) case 1 show ?case using invAM' as invAB' * by presburger case 2 show ?case using var_match[OF m M _ pref_match1, of a] a atLeast0LessThan * ** unfolding v by (metis lessThan_iff) qed next assume matched: "\ ?not_matched" then obtain a' where a'eq: "B(match A a) = Some a'" by auto have a': "a' \ ?M \ match A a' = match A a" using a'eq invAB by (metis ranI) hence "a' < n" "a \ a'" using a M atLeast0LessThan by auto show ?ELSE unfolding a'eq option.sel proof (goal_cases) have inj_dom: "inj_on B (dom B)" by (metis (mono_tags) domD inj_onI invAB) case 1 show ?case (is "(?pref \ ?THEN) \ (\ ?pref \ ?ELSE)") proof (rule; rule) assume ?pref show ?THEN proof (simp only: mem_Collect_eq prod.case, rule conjI, goal_cases) have *: "{ {a} = {b x. b \ match A a \ B b = Some x \ a'\ x" using invAB a' by blast have invAB': "invAB (A[a' := A ! a' + 1]) (B(match A a := Some a)) ({?pref\] invAB' unfolding * using a' as by simp case 2 have "card({a' < n\ a atLeast0LessThan by (metis Suc_eq_plus1 lessThan_iff var_def) qed next assume "\ ?pref" show ?ELSE proof (simp only: mem_Collect_eq prod.case, rule conjI, goal_cases) case 1 have "invAB (A[a := A ! a + 1]) B ?M" using invAB a by (metis match_def nth_list_update_neq ranI) thus ?case using invAM_next[OF invAM a a' \\ ?pref\] using "2" by presburger case 2 show ?case using a v var_next[OF m M _ pref_match1, of a] by (metis Suc_eq_plus1 atLeast0LessThan lessThan_iff) qed qed qed qed qed subsection \Algorithm 5: Data refinement of \B\\ -definition "\ B M = (\b. if b < n \ M!b then Some(B!b) else None)" +definition "\ B N = (\b. if b < n \ N!b then Some(B!b) else None)" -lemma \_Some[simp]: "\ B M b = Some a \ b < n \ M!b \ a = B!b" +lemma \_Some[simp]: "\ B N b = Some a \ b < n \ N!b \ a = B!b" by(auto simp add: \_def) -lemma \update1: "\ \ M ! b; b < length B; b < length M; n = length M \ - \ ran(\ (B[b := a]) (M[b := True])) = ran(\ B M) \ {a}" +lemma \update1: "\ \ N ! b; b < length B; b < length N; n = length N \ + \ ran(\ (B[b := a]) (N[b := True])) = ran(\ B N) \ {a}" by(force simp add: \_def ran_def nth_list_update) -lemma \update2: "\ M!b; b < length B; b < length M; length M = n \ - \ \ (B[b := a]) M = (\ B M)(b := Some a)" +lemma \update2: "\ N!b; b < length B; b < length N; length N = n \ + \ \ (B[b := a]) N = (\ B N)(b := Some a)" by(force simp add: \_def nth_list_update) abbreviation invAB2 :: "nat list \ nat list \ bool list \ nat set \ bool" where -"invAB2 A B M M' == (invAB A (\ B M) M' \ (length B = n \ length M = n))" +"invAB2 A B N M == (invAB A (\ B N) M \ (length B = n \ length N = n))" definition invar1 where -[simp]: "invar1 A B M ai = (invAM A { invAB2 A B M { ai \ n)" +[simp]: "invar1 A B N ai = (invAM A { invAB2 A B N { ai \ n)" definition invar2 where -[simp]: "invar2 A B M ai a \ (invAM A ({ invAB2 A B M ({ a \ ai \ ai < n)" +[simp]: "invar2 A B N ai a \ (invAM A ({ invAB2 A B N ({ a \ ai \ ai < n)" text \First, the `old' version with the more complicated inner variant:\ lemma Gale_Shapley5: -"VARS A B M ai a a' - [ai = 0 \ A = replicate n 0 \ length B = n \ M = replicate n False] +"VARS A B N ai a a' + [ai = 0 \ A = replicate n 0 \ length B = n \ N = replicate n False] WHILE ai < n - INV { invar1 A B M ai } + INV { invar1 A B N ai } VAR { z = n - ai} DO a := ai; - WHILE M ! match A a - INV { invar2 A B M ai a \ z = n-ai } + WHILE N ! match A a + INV { invar2 A B N ai a \ z = n-ai } VAR {var A { a < a' THEN B[match A a] := a; A[a'] := A!a'+1; a := a' ELSE A[a] := A!a+1 FI OD; - B[match A a] := a; M[match A a] := True; ai := ai+1 - OD - [matching A { stable A { opti\<^sub>a A]" -proof (vcg_tc, goal_cases) - case 1 thus ?case (* outer invar holds initially *) - by(auto simp: pref_match_def P_set card_distinct match_def index_nth_id prefers_def opti\<^sub>a_def \_def cong: conj_cong) -next - case 2 (* outer invar and b implies inner invar *) - thus ?case by (auto simp: atLeastLessThanSuc_atLeastAtMost simp flip: atLeastLessThan_eq_atLeastAtMost_diff) -next - case (4 z A B M ai a) (* inner invar' and not b implies outer invar *) - note inv = 4[THEN conjunct1, unfolded invar2_def] - note invAM = inv[THEN conjunct1,THEN conjunct1] - note aai = inv[THEN conjunct1, THEN conjunct2, THEN conjunct2] - show ?case - proof (simp only: mem_Collect_eq prod.case, rule conjI, goal_cases) - case 1 - have *: "{ match A a \ match A ` ({ match A ` ({update1 match_less_n insert_absorb nth_list_update) - next - case 2 thus ?case using 4 by auto - qed -next - case 5 (* outer invar and not b implies post *) - thus ?case using pref_match_stable unfolding invAM_def invar1_def by(metis le_neq_implies_less) -next - case (3 z v A B M ai a) (* preservation of inner invar *) - let ?M = "{ { ai \ ai < n" and invAB: "invAB2 A B M ?M" - and v: "var A { a \ ?M" using invar by auto - have a': "?a \ ?M \ match A ?a = match A a" - using invAB match_less_n[OF A] a matched by (metis \_Some ranI) - have "?M \ { {a} = { B M) (dom (\ B M))" by (metis (mono_tags) domD inj_onI invAB) - have invAB': "invAB (A[B ! match A a := A ! ?a + 1]) (\ (B[match A a := a]) M) ({update2) - show ?case using invAM_swap[OF invAM a a' unstab] invAB' invar a' - unfolding * by (simp add: insert_absorb \update2) - - case 2 - show ?case using v var_next[OF m M \?M \ { pref_match1 \?a < n\] card - by (metis var_def Suc_eq_plus1) - next - case stab: 3 - have *: "\b. b < n \ M!b \ a \ B!b" by (metis invAB ranI \_Some a) - show ?case using invAM_next[OF invAM a a' stab] invar * by (simp add: match_def) - - case 4 - show ?case using v var_next[OF m M \?M \ { pref_match1, of a] a card - by (metis Suc_eq_plus1 var_def) - qed - qed -qed - -text \The definitive version with variant @{const var2}:\ - -lemma Gale_Shapley5_var2: -"VARS A B M ai a a' - [ai = 0 \ A = replicate n 0 \ length B = n \ M = replicate n False] - WHILE ai < n - INV { invar1 A B M ai } - VAR { z = n - ai} - DO a := ai; - WHILE M ! match A a - INV { invar2 A B M ai a \ z = n-ai } - VAR {var2 A} - DO a' := B ! match A a; - IF Q ! match A a' \ a < a' - THEN B[match A a] := a; A[a'] := A!a'+1; a := a' - ELSE A[a] := A!a+1 - FI - OD; - B[match A a] := a; M[match A a] := True; ai := ai+1 + B[match A a] := a; N[match A a] := True; ai := ai+1 OD [matching A { stable A { opti\<^sub>a A]" proof (vcg_tc, goal_cases) case 1 thus ?case (* outer invar holds initially *) by(auto simp: pref_match_def P_set card_distinct match_def index_nth_id prefers_def opti\<^sub>a_def \_def cong: conj_cong) next case 2 (* outer invar and b implies inner invar *) thus ?case by (auto simp: atLeastLessThanSuc_atLeastAtMost simp flip: atLeastLessThan_eq_atLeastAtMost_diff) next case (4 z A B M ai a) (* inner invar' and not b implies outer invar *) note inv = 4[THEN conjunct1, unfolded invar2_def] note invAM = inv[THEN conjunct1,THEN conjunct1] note aai = inv[THEN conjunct1, THEN conjunct2, THEN conjunct2] show ?case proof (simp only: mem_Collect_eq prod.case, rule conjI, goal_cases) case 1 have *: "{ match A a \ match A ` ({ match A ` ({update1 match_less_n insert_absorb nth_list_update) next case 2 thus ?case using 4 by auto qed next case 5 (* outer invar and not b implies post *) thus ?case using pref_match_stable unfolding invAM_def invar1_def by(metis le_neq_implies_less) next - case (3 z v A B M ai a) (* preservation of inner invar *) + case (3 z v A B N ai a) (* preservation of inner invar *) let ?M = "{ { ai \ ai < n" and invAB: "invAB2 A B M ?M" + and matched: "N ! match A a" and as: "a \ ai \ ai < n" and invAB: "invAB2 A B N ?M" + and v: "var A { a \ ?M" using invar by auto + have a': "?a \ ?M \ match A ?a = match A a" + using invAB match_less_n[OF A] a matched by (metis \_Some ranI) + have "?M \ { {a} = { B N) (dom (\ B N))" by (metis (mono_tags) domD inj_onI invAB) + have invAB': "invAB (A[B ! match A a := A ! ?a + 1]) (\ (B[match A a := a]) N) ({update2) + show ?case using invAM_swap[OF invAM a a' unstab] invAB' invar a' + unfolding * by (simp add: insert_absorb \update2) + + case 2 + show ?case using v var_next[OF m M \?M \ { pref_match1 \?a < n\] card + by (metis var_def Suc_eq_plus1) + next + case stab: 3 + have *: "\b. b < n \ N!b \ a \ B!b" by (metis invAB ranI \_Some a) + show ?case using invAM_next[OF invAM a a' stab] invar * by (simp add: match_def) + + case 4 + show ?case using v var_next[OF m M \?M \ { pref_match1, of a] a card + by (metis Suc_eq_plus1 var_def) + qed + qed +qed + +text \The definitive version with variant @{const var2}:\ + +lemma Gale_Shapley5_var2: +"VARS A B N ai a a' + [ai = 0 \ A = replicate n 0 \ length B = n \ N = replicate n False] + WHILE ai < n + INV { invar1 A B N ai } + VAR { z = n - ai} + DO a := ai; + WHILE N ! match A a + INV { invar2 A B N ai a \ z = n-ai } + VAR {var2 A} + DO a' := B ! match A a; + IF Q ! match A a' \ a < a' + THEN B[match A a] := a; A[a'] := A!a'+1; a := a' + ELSE A[a] := A!a+1 + FI + OD; + B[match A a] := a; N[match A a] := True; ai := ai+1 + OD + [matching A { stable A { opti\<^sub>a A]" +proof (vcg_tc, goal_cases) + case 1 thus ?case (* outer invar holds initially *) + by(auto simp: pref_match_def P_set card_distinct match_def index_nth_id prefers_def opti\<^sub>a_def \_def cong: conj_cong) +next + case 2 (* outer invar and b implies inner invar *) + thus ?case by (auto simp: atLeastLessThanSuc_atLeastAtMost simp flip: atLeastLessThan_eq_atLeastAtMost_diff) +next + case (4 z A B N ai a) (* inner invar' and not b implies outer invar *) + note inv = 4[THEN conjunct1, unfolded invar2_def] + note invAM = inv[THEN conjunct1,THEN conjunct1] + note aai = inv[THEN conjunct1, THEN conjunct2, THEN conjunct2] + show ?case + proof (simp only: mem_Collect_eq prod.case, rule conjI, goal_cases) + case 1 + have *: "{ match A a \ match A ` ({ match A ` ({update1 match_less_n insert_absorb nth_list_update) + next + case 2 thus ?case using 4 by auto + qed +next + case 5 (* outer invar and not b implies post *) + thus ?case using pref_match_stable unfolding invAM_def invar1_def by(metis le_neq_implies_less) +next + case (3 z v A B N ai a) (* preservation of inner invar *) + let ?M = "{ { ai \ ai < n" and invAB: "invAB2 A B N ?M" and v: "var2 A = v" using 3 by auto note invar = 3[THEN conjunct1, THEN conjunct1, unfolded invar2_def] let ?a = "B ! match A a" have a: "a < n \ a \ ?M" using invar by auto have a': "?a \ ?M \ match A ?a = match A a" using invAB match_less_n[OF A] a matched by (metis \_Some ranI) have "?M \ { {a} = { B M) (dom (\ B M))" by (metis (mono_tags) domD inj_onI invAB) - have invAB': "invAB (A[B ! match A a := A ! ?a + 1]) (\ (B[match A a := a]) M) ({ B N) (dom (\ B N))" by (metis (mono_tags) domD inj_onI invAB) + have invAB': "invAB (A[B ! match A a := A ! ?a + 1]) (\ (B[match A a := a]) N) ({update2) show ?case using invAM' invAB' invar a' unfolding * by (simp add: insert_absorb \update2) (* case 2 show ?case using v var2_next[OF invAM'] \a' < n\ * atLeast0LessThan by auto*) case 2 show ?case using v var2_next[OF invAM'] * M \?a < n\ a' by (metis subset_Diff_insert) next case stab: 3 note invAM' = invAM_next[OF invAM a a' stab] - have "\b. b < n \ M!b \ a \ B!b" by (metis invAB ranI \_Some a) + have "\b. b < n \ N!b \ a \ B!b" by (metis invAB ranI \_Some a) thus ?case using invAM' invar by (simp add: match_def) case 4 show ?case using v var2_next[OF invAM'] a \?M \ { by blast qed qed qed subsubsection \Algorithm 5.1: single-loop variant\ definition invar2' where -[simp]: "invar2' A B M ai a \ (invAM A ({ invAB2 A B M ({ a \ ai \ ai \ n)" +[simp]: "invar2' A B N ai a \ (invAM A ({ invAB2 A B N ({ a \ ai \ ai \ n)" lemma pres2': -assumes "invar2' A B M ai a" "ai < n" "var A ({ M ! b \ - invar2' A (B[b := a]) (M[b := True]) (ai + 1) (ai + 1) \ var A ({ - (M ! b \ - (Q ! match A a' \ a < a' \ invar2' A1 (B[b := a]) M ai a' \ var A1 ({ - (\ Q ! match A a' \ a < a' \ invar2' A2 B M ai a \ var A2 ({ N ! b \ + invar2' A (B[b := a]) (N[b := True]) (ai + 1) (ai + 1) \ var A ({ + (N ! b \ + (Q ! match A a' \ a < a' \ invar2' A1 (B[b := a]) N ai a' \ var A1 ({ + (\ Q ! match A a' \ a < a' \ invar2' A2 B N ai a \ var A2 ({ { ai \ ai < n" and invAB: "invAB2 A B M ?M" + and v: "var A ?M = v" and as: "a \ ai \ ai < n" and invAB: "invAB2 A B N ?M" using assms by auto note invar = assms(1) note pref_match1 = preferred_subset_match_if_invAM[OF invAM] have a: "a < n \ a \ ?M" using as by (simp) show ?thesis (is "(\ ?matched \ ?THEN) \ (?matched \ ?ELSE)") proof (rule; rule) assume "\ ?matched" then have nm: "match A a \ match A ` ?M" using invAB unfolding ran_def apply (clarsimp simp: set_eq_iff) by metis show ?THEN proof(rule conjI, goal_cases) have *: "{ {a} = {\ ?matched\ ** by (simp add: A a \update1 match_less_n nth_list_update) case 1 show ?case using invAM' as invAB' * by (simp add: Suc_le_eq plus_1_eq_Suc) case 2 show ?case using var_match[OF m M _ pref_match1, of a] a atLeast0LessThan * ** unfolding v by (metis lessThan_iff) qed next assume matched: "?matched" let ?a = "B ! match A a" have a': "?a \ ?M \ match A ?a = match A a" using invAB match_less_n[OF A] a matched after by (metis \_Some ranI) hence "?a < n" "a \ ?a" using a M atLeast0LessThan by auto - have inj_dom: "inj_on (\ B M) (dom (\ B M))" by (metis (mono_tags) domD inj_onI invAB) + have inj_dom: "inj_on (\ B N) (dom (\ B N))" by (metis (mono_tags) domD inj_onI invAB) show ?ELSE (is "(?pref \ ?THEN) \ (\ ?pref \ ?ELSE)") proof (rule; rule) assume ?pref show ?THEN proof (rule conjI, goal_cases) have *: "{ {a} = {b match A a \ M!b \ ?a \ B!b" + have a'neq: "\b match A a \ N!b \ ?a \ B!b" using invAB a' by force - have invAB': "invAB (A[B ! match A a := A ! ?a + 1]) (\ (B[match A a := a]) M) ({ (B[match A a := a]) N) ({update2) case 1 show ?case using invAM_swap[OF invAM a a'] \?pref\ invAB invAB' unfolding * using a' as by simp case 2 have "card({?a < n\ a atLeast0LessThan by (metis Suc_eq_plus1 lessThan_iff var_def after) qed next assume "\ ?pref" show ?ELSE proof (rule conjI, goal_cases) case 1 - have "invAB2 (A[a := A ! a + 1]) B M ?M" using invAB a + have "invAB2 (A[a := A ! a + 1]) B N ?M" using invAB a by (metis match_def nth_list_update_neq ranI) thus ?case using invAM_next[OF invAM a a'] \\ ?pref\ invar after by (meson invar2'_def) case 2 show ?case using a v var_next[OF m M _ pref_match1, of a] after by (metis Suc_eq_plus1 atLeast0LessThan lessThan_iff) qed qed qed qed -lemma Gale_Shapley5_1: "VARS A B M a a' ai b - [ai = 0 \ a = 0 \ A = replicate n 0 \ length B = n \ M = replicate n False] +lemma Gale_Shapley5_1: "VARS A B N a a' ai b + [ai = 0 \ a = 0 \ A = replicate n 0 \ length B = n \ N = replicate n False] WHILE ai < n - INV { invar2' A B M ai a } + INV { invar2' A B N ai a } VAR {var A ({ M ! b - THEN B[b] := a; M[b] := True; ai := ai + 1; a := ai + IF \ N ! b + THEN B[b] := a; N[b] := True; ai := ai + 1; a := ai ELSE a' := B ! b; IF Q ! match A a' \ a < a' THEN B[b] := a; A[a'] := A!a'+1; a := a' ELSE A[a] := A!a+1 FI FI OD [matching A { stable A { opti\<^sub>a A]" proof (vcg_tc, goal_cases) case 1 thus ?case by(auto simp: pref_match_def P_set card_distinct match_def index_nth_id prefers_def opti\<^sub>a_def \_def cong: conj_cong) next case 3 thus ?case using pref_match_stable using atLeast0_lessThan_Suc by force next - case (2 v A B M a a' ai b) + case (2 v A B N a a' ai b) thus ?case using pres2' by (simp only:mem_Collect_eq prod.case) blast qed subsection \Algorithm 6: replace \Q\ by ranking \R\\ lemma inner_to_outer: -assumes inv: "invar2 A B M ai a \ b = match A a" and not_b: "\ M ! b" -shows "invar1 A (B[b := a]) (M[b := True]) (ai+1)" +assumes inv: "invar2 A B N ai a \ b = match A a" and not_b: "\ N ! b" +shows "invar1 A (B[b := a]) (N[b := True]) (ai+1)" proof - note invAM = inv[unfolded invar2_def, THEN conjunct1,THEN conjunct1] have *: "{ match A a \ match A ` ({ match A ` ({update1 insert_absorb nth_list_update) qed lemma inner_pres: assumes R: "\ba1a2 Q ! b \ a1 < a2" and - inv: "invar2 A B M ai a" and m: "M ! b" and v: "var A { invar2 A1 (B[b:=a]) M ai a' \ var A1 { - (\ r ! a < r ! a' \ invar2 A2 B M ai a \ var A2 { invar2 A1 (B[b:=a]) N ai a' \ var A1 { + (\ r ! a < r ! a' \ invar2 A2 B N ai a \ var A2 { { B M) ?M" + have A: "wf A" and M: "?M \ { B N) ?M" and mat: "matching A ?M" and as: "a \ ai \ ai < n" using inv' by auto note pref_match1 = preferred_subset_match_if_invAM[OF invAM] let ?a = "B ! match A a" have a: "a < n \ a \ ?M" using inv by auto have a': "?a \ ?M \ match A ?a = match A a" using invAB match_less_n[OF A] a m inv by (metis \_Some ranI \b = _\) have "?M \ { {a} = { a < a'" using R a a' as Q_set P_set match_less_n[OF A] n_def length_Q R by (simp) - have inj_dom: "inj_on (\ B M) (dom (\ B M))" by (metis (mono_tags) domD inj_onI invAB) - have invAB': "invAB A1 (\ (B[match A a := a]) M) ({ B N) (dom (\ B N))" by (metis (mono_tags) domD inj_onI invAB) + have invAB': "invAB A1 (\ (B[match A a := a]) N) ({update2 inv') show ?case using invAM_swap[OF invAM a a'] unstab invAB' inv a' unfolding * by (simp) next case 2 show ?case using v var_next[OF mat M \?M \ { pref_match1 \?a < n\] card assms(5,7,9) by (metis Suc_eq_plus1 var_def) next - have *: "\b. b < n \ M!b \ a \ B!b" by (metis invAB ranI \_Some a) + have *: "\b. b < n \ N!b \ a \ B!b" by (metis invAB ranI \_Some a) case 3 hence unstab: "\ Q ! match A a' \ a < a'" using R a a' as Q_set P_set match_less_n[OF A] n_def length_Q by (simp add: ranking_iff_pref) then show ?case using invAM_next[OF invAM a a'] 3 inv * by (simp add: match_def) next case 4 show ?case using v var_next[OF mat M \?M \ { pref_match1, of a] a card assms(6) by (metis Suc_eq_plus1 var_def) qed qed text \First, the `old' version with the more complicated inner variant:\ lemma Gale_Shapley6: assumes "R = map ranking Q" shows -"VARS A B M ai a a' b r - [ai = 0 \ A = replicate n 0 \ length B = n \ M = replicate n False] +"VARS A B N ai a a' b r + [ai = 0 \ A = replicate n 0 \ length B = n \ N = replicate n False] WHILE ai < n - INV { invar1 A B M ai } + INV { invar1 A B N ai } VAR {z = n - ai} DO a := ai; b := match A a; - WHILE M ! b - INV { invar2 A B M ai a \ b = match A a \ z = n-ai } + WHILE N ! b + INV { invar2 A B N ai a \ b = match A a \ z = n-ai } VAR {var A { stable A { opti\<^sub>a A]" proof (vcg_tc, goal_cases) case 1 thus ?case (* outer invar holds initially *) by(auto simp: stable_def pref_match_def P_set card_distinct match_def index_nth_id prefers_def opti\<^sub>a_def \_def cong: conj_cong) next case 2 (* outer invar and b implies inner invar *) thus ?case by (auto simp: atLeastLessThanSuc_atLeastAtMost simp flip: atLeastLessThan_eq_atLeastAtMost_diff) next case 3 (* preservation of inner invar *) have R: "\ba1a2 Q ! b \ a1 < a2" by (simp add: Q_set \R = _\ length_Q ranking_iff_pref) show ?case proof (simp only: mem_Collect_eq prod.case, goal_cases) case 1 show ?case using inner_pres[OF R _ _ refl refl refl] 3 by blast qed next case 4 (* inner invar' and not b implies outer invar *) show ?case proof (simp only: mem_Collect_eq prod.case, rule conjI, goal_cases) case 1 show ?case using 4 inner_to_outer by blast next case 2 thus ?case using 4 by auto qed next case 5 (* outer invar and not b implies post *) thus ?case using pref_match_stable unfolding invAM_def invar1_def by(metis le_neq_implies_less) qed lemma inner_pres_var2: assumes R: "\ba1a2 Q ! b \ a1 < a2" and - inv: "invar2 A B M ai a" and m: "M ! b" and v: "var2 A = v" + inv: "invar2 A B N ai a" and m: "N ! b" and v: "var2 A = v" and after: "A1 = A[a' := A ! a' + 1]" "A2 = A[a := A ! a + 1]" "a' = B!b" "r = R ! match A a'" "b = match A a" -shows "(r ! a < r ! a' \ invar2 A1 (B[b:=a]) M ai a' \ var2 A1 < v) \ - (\ r ! a < r ! a' \ invar2 A2 B M ai a \ var2 A2 < v)" +shows "(r ! a < r ! a' \ invar2 A1 (B[b:=a]) N ai a' \ var2 A1 < v) \ + (\ r ! a < r ! a' \ invar2 A2 B N ai a \ var2 A2 < v)" proof - let ?M = "{ { B M) ?M" + have A: "wf A" and M: "?M \ { B N) ?M" and mat: "matching A ?M" and as: "a \ ai \ ai < n" using inv' by auto let ?a = "B ! match A a" have a: "a < n \ a \ ?M" using inv by auto have a': "?a \ ?M \ match A ?a = match A a" using invAB match_less_n[OF A] a m inv by (metis \_Some ranI \b = _\) have "?M \ { {a} = { a < a'" using R a a' as Q_set P_set match_less_n[OF A] n_def length_Q R by (simp) - have inj_dom: "inj_on (\ B M) (dom (\ B M))" by (metis (mono_tags) domD inj_onI invAB) - have invAB': "invAB A1 (\ (B[match A a := a]) M) ({ B N) (dom (\ B N))" by (metis (mono_tags) domD inj_onI invAB) + have invAB': "invAB A1 (\ (B[match A a := a]) N) ({update2 inv') show ?case using invAM' unstab invAB' inv a' unfolding * by (simp) case 2 show ?case using v var2_next[OF invAM'] assms(5,7,9) * M \?a < n\ a' by (metis subset_Diff_insert unstab) next - have *: "\b. b < n \ M!b \ a \ B!b" by (metis invAB ranI \_Some a) + have *: "\b. b < n \ N!b \ a \ B!b" by (metis invAB ranI \_Some a) note invAM' = invAM_next[OF invAM a a'] case 3 hence unstab: "\ Q ! match A a' \ a < a'" using R a a' as Q_set P_set match_less_n[OF A] n_def length_Q by (simp add: ranking_iff_pref) then show ?case using invAM' 3 inv * by (simp add: match_def) case 4 show ?case using v var2_next[OF invAM'] a assms(6,7,9) \?M \ { unstab by fastforce qed qed text \The definitive version with variant @{const var2}:\ lemma Gale_Shapley6_var2: assumes "R = map ranking Q" shows -"VARS A B M ai a a' b r - [ai = 0 \ A = replicate n 0 \ length B = n \ M = replicate n False] +"VARS A B N ai a a' b r + [ai = 0 \ A = replicate n 0 \ length B = n \ N = replicate n False] WHILE ai < n - INV { invar1 A B M ai } + INV { invar1 A B N ai } VAR {z = n - ai} DO a := ai; b := match A a; - WHILE M ! b - INV { invar2 A B M ai a \ b = match A a \ z = n-ai } + WHILE N ! b + INV { invar2 A B N ai a \ b = match A a \ z = n-ai } VAR {var2 A} DO a' := B ! b; r := R ! match A a'; IF r ! a < r ! a' THEN B[b] := a; A[a'] := A!a'+1; a := a' ELSE A[a] := A!a+1 FI; b := match A a OD; - B[b] := a; M[b] := True; ai := ai+1 + B[b] := a; N[b] := True; ai := ai+1 OD [matching A { stable A { opti\<^sub>a A]" proof (vcg_tc, goal_cases) case 1 thus ?case (* outer invar holds initially *) by(auto simp: stable_def pref_match_def P_set card_distinct match_def index_nth_id prefers_def opti\<^sub>a_def \_def cong: conj_cong) next case 2 (* outer invar and b implies inner invar *) thus ?case by (auto simp: atLeastLessThanSuc_atLeastAtMost simp flip: atLeastLessThan_eq_atLeastAtMost_diff) next case 3 (* preservation of inner invar *) have R: "\ba1a2 Q ! b \ a1 < a2" by (simp add: Q_set \R = _\ length_Q ranking_iff_pref) show ?case proof (simp only: mem_Collect_eq prod.case, goal_cases) case 1 show ?case using inner_pres_var2[OF R _ _ refl refl refl] 3 by blast qed next case 4 (* inner invar' and not b implies outer invar *) show ?case proof (simp only: mem_Collect_eq prod.case, rule conjI, goal_cases) case 1 show ?case using 4 inner_to_outer by blast next case 2 thus ?case using 4 by auto qed next case 5 (* outer invar and not b implies post *) thus ?case using pref_match_stable unfolding invAM_def invar1_def by(metis le_neq_implies_less) qed text \A less precise version where the inner variant does not depend on variables changed in the outer loop. Thus the inner variant is an upper bound on the number of executions of the inner loop test/body. Superseded by the @{const var2} version.\ lemma var0_next2: assumes "wf (A[a' := A ! a' + 1])" "a' < n" shows "var0 (A[a' := A ! a' + 1]) {a n^2" using sumA_ub[OF assms(1)] 0 by (simp add: power2_eq_square algebra_simps le_diff_conv2) have "(\aaba1a2 Q ! b \ a1 < a2" and - inv: "invar2 A B M ai a" and m: "M ! b" and v: "var0 A { invar2 A1 (B[b:=a]) M ai a' \ var0 A1 { - (\ r ! a < r ! a' \ invar2 A2 B M ai a \ var0 A2 { invar2 A1 (B[b:=a]) N ai a' \ var0 A1 { + (\ r ! a < r ! a' \ invar2 A2 B N ai a \ var0 A2 { { B M) ?M" + have A: "wf A" and M: "?M \ { B N) ?M" and mat: "matching A ?M" and as: "a \ ai \ ai < n" using inv' by auto note pref_match1 = preferred_subset_match_if_invAM[OF invAM] let ?a = "B ! match A a" have a: "a < n \ a \ ?M" using inv by auto have a': "?a \ ?M \ match A ?a = match A a" using invAB match_less_n[OF A] a m inv by (metis \_Some ranI \b = _\) have "?M \ { {a} = { a < a'" using R a a' as Q_set P_set match_less_n[OF A] n_def length_Q R by (simp) - have inj_dom: "inj_on (\ B M) (dom (\ B M))" by (metis (mono_tags) domD inj_onI invAB) - have invAB': "invAB A1 (\ (B[match A a := a]) M) ({ B N) (dom (\ B N))" by (metis (mono_tags) domD inj_onI invAB) + have invAB': "invAB A1 (\ (B[match A a := a]) N) ({update2 inv') show ?case using invAM_swap[OF invAM a a'] unstab invAB' inv a' unfolding * by (simp add: insert_absorb \update2) next case 2 hence unstab: "Q ! match A a' \ a < a'" using R a a' as Q_set P_set match_less_n[OF A] n_def length_Q R by (simp) from invAM_swap[OF invAM a a'] unstab have wf: "wf (A[a' := A ! a' + 1])" by auto show ?case using v var0_next2[OF wf] using \B ! match A a < n\ assms(5,7,9) by blast next - have *: "\b. b < n \ M!b \ a \ B!b" by (metis invAB ranI \_Some a) + have *: "\b. b < n \ N!b \ a \ B!b" by (metis invAB ranI \_Some a) case 3 hence unstab: "\ Q ! match A a' \ a < a'" using R a a' as Q_set P_set match_less_n[OF A] n_def length_Q by (simp add: ranking_iff_pref) then show ?case using invAM_next[OF invAM a a'] 3 inv * by (simp add: match_def) next case 4 hence unstab: "\ Q ! match A a' \ a < a'" using R a a' as Q_set P_set match_less_n[OF A] n_def length_Q by (simp add: ranking_iff_pref) from invAM_next[OF invAM a a'] unstab have wf: "wf (A[a := A ! a + 1])" by auto show ?case using v var0_next2[OF wf] a using assms(6) by presburger qed qed lemma Gale_Shapley6': assumes "R = map ranking Q" shows -"VARS A B M ai a a' b r - [ai = 0 \ A = replicate n 0 \ length B = n \ M = replicate n False] +"VARS A B N ai a a' b r + [ai = 0 \ A = replicate n 0 \ length B = n \ N = replicate n False] WHILE ai < n - INV { invar1 A B M ai } + INV { invar1 A B N ai } VAR {z = n - ai} DO a := ai; b := match A a; - WHILE M ! b - INV { invar2 A B M ai a \ b = match A a \ z = n-ai } + WHILE N ! b + INV { invar2 A B N ai a \ b = match A a \ z = n-ai } VAR {var0 A { stable A { opti\<^sub>a A]" proof (vcg_tc, goal_cases) case 1 thus ?case (* outer invar holds initially *) by(auto simp: stable_def pref_match_def P_set card_distinct match_def index_nth_id prefers_def opti\<^sub>a_def \_def cong: conj_cong) next case 2 (* outer invar and b implies inner invar *) thus ?case by (auto simp: atLeastLessThanSuc_atLeastAtMost simp flip: atLeastLessThan_eq_atLeastAtMost_diff) next case 3 (* preservation of inner invar *) have R: "\ba1a2 Q ! b \ a1 < a2" by (simp add: Q_set \R = _\ length_Q ranking_iff_pref) show ?case proof (simp only: mem_Collect_eq prod.case, goal_cases) case 1 show ?case using inner_pres2[OF R _ _ refl refl refl] 3 by blast qed next case 4 (* inner invar' and not b implies outer invar *) show ?case proof (simp only: mem_Collect_eq prod.case, rule conjI, goal_cases) case 1 show ?case using 4 inner_to_outer by blast next case 2 thus ?case using 4 by auto qed next case 5 (* outer invar and not b implies post *) thus ?case using pref_match_stable unfolding invAM_def invar1_def by(metis le_neq_implies_less) qed subsubsection \Algorithm 6.1: single-loop variant\ lemma R_iff_P: -assumes "R = map ranking Q" "invar2' A B M ai a" "ai < n" "M ! match A a" +assumes "R = map ranking Q" "invar2' A B N ai a" "ai < n" "N ! match A a" shows "(R ! match A (B ! match A a) ! a < R ! match A (B ! match A a) ! (B ! match A a)) = (Q ! match A (B ! match A a) \ a < B ! match A a)" proof - have R: "\ba1a2 Q ! b \ a1 < a2" by (simp add: Q_set \R = _\ length_Q ranking_iff_pref) let ?M = "{ { { ?M" - using invAB match_less_n[OF A] as \M!match A a\ by (metis \_Some ranI) + using invAB match_less_n[OF A] as \N!match A a\ by (metis \_Some ranI) hence "B ! match A a < n" using M by auto thus ?thesis using assms R match_less_n by auto qed lemma Gale_Shapley6_1: assumes "R = map ranking Q" -shows "VARS A B M a a' ai b r - [ai = 0 \ a = 0 \ A = replicate n 0 \ length B = n \ M = replicate n False] +shows "VARS A B N a a' ai b r + [ai = 0 \ a = 0 \ A = replicate n 0 \ length B = n \ N = replicate n False] WHILE ai < n - INV { invar2' A B M ai a } + INV { invar2' A B N ai a } VAR {var A ({ M ! b - THEN B[b] := a; M[b] := True; ai := ai + 1; a := ai + IF \ N ! b + THEN B[b] := a; N[b] := True; ai := ai + 1; a := ai ELSE a' := B ! b; r := R ! match A a'; IF r ! a < r ! a' THEN B[b] := a; A[a'] := A!a'+1; a := a' ELSE A[a] := A!a+1 FI FI OD [matching A { stable A { opti\<^sub>a A]" proof (vcg_tc, goal_cases) case 1 thus ?case by(auto simp: pref_match_def P_set card_distinct match_def index_nth_id prefers_def opti\<^sub>a_def \_def cong: conj_cong) next case 3 thus ?case using pref_match_stable atLeast0_lessThan_Suc by force next - case (2 v A B M a a' ai) - have R': "M ! match A a \ + case (2 v A B N a a' ai) + have R': "N ! match A a \ (R ! match A (B ! match A a) ! a < R ! match A (B ! match A a) ! (B ! match A a)) = (Q ! match A (B ! match A a) \ a < B ! match A a)" using R_iff_P 2 assms by blast show ?case apply(simp only:mem_Collect_eq prod.case) - using 2 R' pres2'[of A B M ai a] by presburger + using 2 R' pres2'[of A B N ai a] by presburger qed (* TODO: rm? *) lemma Gale_Shapley6_1_explicit: assumes "R = map ranking Q" -shows "VARS A B M a a' ai b r - [ai = 0 \ a = 0 \ A = replicate n 0 \ length B = n \ M = replicate n False] +shows "VARS A B N a a' ai b r + [ai = 0 \ a = 0 \ A = replicate n 0 \ length B = n \ N = replicate n False] WHILE ai < n - INV { invar2' A B M ai a } + INV { invar2' A B N ai a } VAR {var A ({ M ! b - THEN B[b] := a; M[b] := True; ai := ai + 1; a := ai + IF \ N ! b + THEN B[b] := a; N[b] := True; ai := ai + 1; a := ai ELSE a' := B ! b; r := R ! match A a'; IF r ! a < r ! a' THEN B[b] := a; A[a'] := A!a'+1; a := a' ELSE A[a] := A!a+1 FI FI OD [matching A { stable A { opti\<^sub>a A]" proof (vcg_tc, goal_cases) case 1 thus ?case by(auto simp: pref_match_def P_set card_distinct match_def index_nth_id prefers_def opti\<^sub>a_def \_def cong: conj_cong) next case 3 thus ?case using pref_match_stable atLeast0_lessThan_Suc by force next - case (2 v A B M a a' ai b) + case (2 v A B N a a' ai b) let ?M = "{ { ai \ ai < n" and invAB: "invAB2 A B M ?M" + and v: "var A ?M = v" and as: "a \ ai \ ai < n" and invAB: "invAB2 A B N ?M" using 2 by auto note invar = 2[THEN conjunct1,THEN conjunct1] note pref_match' = pref_match[THEN pref_match'_iff[OF A, THEN iffD2]] hence pref_match1: "\a match A ` ?M" unfolding pref_match'_def by blast have a: "a < n \ a \ ?M" using as by (simp) show ?case (is "(?not_matched \ ?THEN) \ (\ ?not_matched \ ?ELSE)") proof (rule; rule) assume ?not_matched then have nm: "match A a \ match A ` ?M" using invAB unfolding ran_def apply (clarsimp simp: set_eq_iff) by metis show ?THEN proof(simp only:mem_Collect_eq prod.case, rule conjI, goal_cases) have *: "{ {a} = {?not_matched\ ** by (simp add: A a \update1 match_less_n nth_list_update) case 1 show ?case using invAM' as invAB' * by (simp add: Suc_le_eq plus_1_eq_Suc) case 2 show ?case using var_match[OF m M _ pref_match1, of a] a atLeast0LessThan * ** unfolding v by (metis lessThan_iff) qed next assume matched: "\ ?not_matched" let ?a = "B ! match A a" have a': "?a \ ?M \ match A ?a = match A a" using invAB match_less_n[OF A] a matched by (metis \_Some ranI) hence "?a < n" "a \ ?a" using a M atLeast0LessThan by auto - have inj_dom: "inj_on (\ B M) (dom (\ B M))" by (metis (mono_tags) domD inj_onI invAB) + have inj_dom: "inj_on (\ B N) (dom (\ B N))" by (metis (mono_tags) domD inj_onI invAB) show ?ELSE (is "(?pref \ ?THEN) \ (\ ?pref \ ?ELSE)") proof (rule; rule) assume ?pref show ?THEN proof (simp only: mem_Collect_eq prod.case, rule conjI, goal_cases) have *: "{ {a} = {b match A a \ M!b \ ?a \ B!b" + have a'neq: "\b match A a \ N!b \ ?a \ B!b" using invAB a' by force - have invAB': "invAB (A[B ! match A a := A ! ?a + 1]) (\ (B[match A a := a]) M) ({ (B[match A a := a]) N) ({update2) have pref: "Q ! match A ?a \ a < ?a" using A Q_set \?a < n\ \?pref\ a assms length_Q by(auto simp: match_less_n ranking_iff_pref) case 1 show ?case (* changed *) using invAM_swap[OF invAM a a' pref] invAB invAB' a' as unfolding * by (simp add: match_less_n ranking_iff_pref) case 2 have "card({?a < n\ a atLeast0LessThan by (metis Suc_eq_plus1 lessThan_iff var_def) qed next assume "\ ?pref" show ?ELSE proof (simp only: mem_Collect_eq prod.case, rule conjI, goal_cases) case 1 - have "invAB2 (A[a := A ! a + 1]) B M ?M" using invAB a + have "invAB2 (A[a := A ! a + 1]) B N ?M" using invAB a by (metis match_def nth_list_update_neq ranI) thus ?case using invAM_next[OF invAM a a'] \\ ?pref\ \B ! match A a < n\ Q_set 2 assms by (simp add: invar2'_def length_Q match_less_n ranking_iff_pref) (* changed *) case 2 show ?case using a v var_next[OF m M _ pref_match1, of a] by (metis Suc_eq_plus1 atLeast0LessThan lessThan_iff) qed qed qed qed end subsection \Functional implementation\ definition -"gs_inner P R M = - while (\(A,B,a,b). M!b) +"gs_inner P R N = + while (\(A,B,a,b). N!b) (\(A,B,a,b). let a' = B ! b; r = R ! (P ! a' ! (A ! a')) in let (A, B, a) = if r ! a < r ! a' then (A[a' := A!a' + 1], B[b := a], a') else (A[a := A!a + 1], B, a) in (A, B, a, P ! a ! (A ! a)))" definition "gs n P R = - while (\(A,B,M,ai). ai < n) - (\(A,B,M,ai). - let (A,B,a,b) = gs_inner P R M (A, B, ai, P ! ai ! (A ! ai)) - in (A, B[b:=a], M[b:=True], ai+1)) + while (\(A,B,N,ai). ai < n) + (\(A,B,N,ai). + let (A,B,a,b) = gs_inner P R N (A, B, ai, P ! ai ! (A ! ai)) + in (A, B[b:=a], N[b:=True], ai+1)) (replicate n 0, replicate n 0, replicate n False,0)" definition "gs1 n P R = - while (\(A,B,M,ai,a). ai < n) - (\(A,B,M,ai,a). + while (\(A,B,N,ai,a). ai < n) + (\(A,B,N,ai,a). let b = P ! a ! (A ! a) in - if \ M ! b - then (A, B[b := a], M[b := True], ai+1, ai+1) + if \ N ! b + then (A, B[b := a], N[b := True], ai+1, ai+1) else let a' = B ! b; r = R ! (P ! a' ! (A ! a')) in if r ! a < r ! a' - then (A[a' := A!a'+1], B[b := a], M, ai, a') - else (A[a := A!a+1], B, M, ai, a)) + then (A[a' := A!a'+1], B[b := a], N, ai, a') + else (A[a := A!a+1], B, N, ai, a)) (replicate n 0, replicate n 0, replicate n False, 0, 0)" context Pref begin lemma gs_inner: assumes "R = map ranking Q" -assumes "invar2 A B M ai a" "b = match A a" -shows "gs_inner P R M (A, B, a, b) = (A',B',a',b') \ invar1 A' (B'[b' := a']) (M[b' := True]) (ai+1)" +assumes "invar2 A B N ai a" "b = match A a" +shows "gs_inner P R N (A, B, a, b) = (A',B',a',b') \ invar1 A' (B'[b' := a']) (N[b' := True]) (ai+1)" unfolding gs_inner_def -proof(rule while_rule2[where P = "\(A,B,a,b). invar2 A B M ai a \ b = match A a" +proof(rule while_rule2[where P = "\(A,B,a,b). invar2 A B N ai a \ b = match A a" and r = "measure (%(A, B, a, b). Pref.var P A {ba1a2 Q ! b \ a1 < a2" by (simp add: Q_set \R = _\ length_Q ranking_iff_pref) show ?case proof(rule, goal_cases) case 1 show ?case using inv apply(simp only: s prod.case Let_def split: if_split) - using inner_pres[OF R _ _ refl refl refl refl refl, of A B M ai a b] + using inner_pres[OF R _ _ refl refl refl refl refl, of A B N ai a b] unfolding invar2_def match_def by presburger case 2 show ?case using inv apply(simp only: s prod.case Let_def in_measure split: if_split) - using inner_pres[OF R _ _ refl refl refl refl refl, of A B M ai a b] + using inner_pres[OF R _ _ refl refl refl refl refl, of A B N ai a b] unfolding invar2_def match_def by presburger qed next case 3 show ?case proof (rule, goal_cases) case 1 show ?case by(rule inner_to_outer[OF 3[unfolded 1 prod.case]]) qed next case 4 show ?case by simp qed lemma gs: assumes "R = map ranking Q" -shows "gs n P R = (A,BMai) \ matching A { stable A { opti\<^sub>a A" +shows "gs n P R = (A,BNai) \ matching A { stable A { opti\<^sub>a A" unfolding gs_def -proof(rule while_rule2[where P = "\(A,B,M,ai). invar1 A B M ai" - and r = "measure(\(A,B,M,ai). n - ai)"], goal_cases) +proof(rule while_rule2[where P = "\(A,B,N,ai). invar1 A B N ai" + and r = "measure(\(A,B,N,ai). n - ai)"], goal_cases) case 1 show ?case by(auto simp: stable_def pref_match_def P_set card_distinct match_def index_nth_id prefers_def opti\<^sub>a_def \_def cong: conj_cong) next case (2 s) - obtain A B M ai where s: "s = (A, B, M, ai)" + obtain A B N ai where s: "s = (A, B, N, ai)" using prod_cases4 by blast - have 1: "invar2 A B M ai ai" using 2 s + have 1: "invar2 A B N ai ai" using 2 s by (auto simp: atLeastLessThanSuc_atLeastAtMost simp flip: atLeastLessThan_eq_atLeastAtMost_diff) show ?case using 2 s gs_inner[OF \R = _ \ 1] by (auto simp: match_def simp del: invar1_def split: prod.split) next case 3 thus ?case using pref_match_stable by auto next case 4 show ?case by simp qed lemma gs1: assumes "R = map ranking Q" -shows "gs1 n P R = (A,BMaia) \ matching A { stable A { opti\<^sub>a A" +shows "gs1 n P R = (A,BNaia) \ matching A { stable A { opti\<^sub>a A" unfolding gs1_def -proof(rule while_rule2[where P = "\(A,B,M,ai,a). invar2' A B M ai a" - and r = "measure (%(A, B, M, ai, a). Pref.var P A ({a_def \_def cong: conj_cong) next case (2 s) - obtain A B M ai a where s: "s = (A, B, M, ai, a)" + obtain A B N ai a where s: "s = (A, B, N, ai, a)" using prod_cases5 by blast - hence 1: "invar2' A B M ai a" "ai < n" using 2 by (simp_all) - have R': "M ! match A a \ + hence 1: "invar2' A B N ai a" "ai < n" using 2 by (simp_all) + have R': "N ! match A a \ (R ! match A (B ! match A a) ! a < R ! match A (B ! match A a) ! (B ! match A a)) = (Q ! match A (B ! match A a) \ a < B ! match A a)" using R_iff_P[OF assms 1] by linarith show ?case using 1 R' pres2'[OF 1] apply(simp only: s mem_Collect_eq prod.case Let_def in_measure match_def split: if_split) by blast next case (3 s) - obtain B M ai a where "BMaia = (B, M, ai, a)" + obtain B N ai a where "BNaia = (B, N, ai, a)" using prod_cases4 by blast with 3 show ?case using pref_match_stable atLeast0_lessThan_Suc by force next case 4 show ?case by simp qed end subsection \Executable functional Code\ definition "Gale_Shapley P Q = (if Pref P Q then Some (fst (gs (length P) P (map ranking Q))) else None)" theorem gs: "\ Pref P Q; n = length P \ \ \A. Gale_Shapley P Q = Some(A) \ Pref.matching P A { Pref.stable P Q A { Pref.opti\<^sub>a P Q A" unfolding Gale_Shapley_def using Pref.gs by (metis fst_conv surj_pair) declare Pref_def [code] definition "Gale_Shapley1 P Q = (if Pref P Q then Some (fst (gs1 (length P) P (map ranking Q))) else None)" theorem gs1: "\ Pref P Q; n = length P \ \ \A. Gale_Shapley1 P Q = Some(A) \ Pref.matching P A { Pref.stable P Q A { Pref.opti\<^sub>a P Q A" unfolding Gale_Shapley1_def using Pref.gs1 by (metis fst_conv surj_pair) declare Pref_def [code] text \Two examples from Gusfield and Irving:\ lemma "Gale_Shapley [[3,0,1,2], [1,2,0,3], [1,3,2,0], [2,0,3,1]] [[3,0,2,1], [0,2,1,3], [0,1,2,3], [3,0,2,1]] = Some[0,1,0,1]" by eval lemma "Gale_Shapley1 [[4,6,0,1,5,7,3,2], [1,2,6,4,3,0,7,5], [7,4,0,3,5,1,2,6], [2,1,6,3,0,5,7,4], [6,1,4,0,2,5,7,3], [0,5,6,4,7,3,1,2], [1,4,6,5,2,3,7,0], [2,7,3,4,6,1,5,0]] [[4,2,6,5,0,1,7,3], [7,5,2,4,6,1,0,3], [0,4,5,1,3,7,6,2], [7,6,2,1,3,0,4,5], [5,3,6,2,7,0,1,4], [1,7,4,2,3,5,6,0], [6,4,1,0,7,5,3,2], [6,3,0,4,1,2,5,7]] = Some [0, 1, 0, 5, 0, 0, 0, 2]" by eval end \ No newline at end of file diff --git a/thys/Gale_Shapley/Gale_Shapley2.thy b/thys/Gale_Shapley/Gale_Shapley2.thy --- a/thys/Gale_Shapley/Gale_Shapley2.thy +++ b/thys/Gale_Shapley/Gale_Shapley2.thy @@ -1,377 +1,377 @@ (* Stepwise refinement of the Gale-Shapley algorithm down to executable code. Author: Tobias Nipkow *) section \Part 2: Refinement from lists to arrays\ theory Gale_Shapley2 imports Gale_Shapley1 "Collections.Diff_Array" begin text \Refinement from lists to arrays, resulting in a linear (in the input size, which is \n^2\) time algorithm.\ abbreviation "array \ new_array" notation array_get (infixl "!!" 100) notation array_set ("_[_ ::= _]" [1000,0,0] 900) abbreviation "list \ list_of_array" lemma list_array: "list (array x n) = replicate n x" by (simp add: new_array_def) lemma array_get: "a !! i = list a ! i" by (cases a) simp context Pref begin subsection \Algorithm 7: Arrays\ definition "match_array A a = P ! a ! (A !! a)" lemma match_array: "match_array A a = match (list A) a" by(cases A) (simp add: match_array_def match_def) lemmas array_abs = match_array array_list_of_set array_get lemma Gale_Shapley7: assumes "R = map ranking Q" shows -"VARS A B M ai a a' b r - [ai = 0 \ A = array 0 n \ B = array 0 n \ M = array False n] +"VARS A B N ai a a' b r + [ai = 0 \ A = array 0 n \ B = array 0 n \ N = array False n] WHILE ai < n - INV { invar1 (list A) (list B) (list M) ai } + INV { invar1 (list A) (list B) (list N) ai } VAR {z = n - ai} DO a := ai; b := match_array A a; - WHILE M !! b - INV { invar2 (list A) (list B) (list M) ai a \ b = match_array A a \ z = n-ai } + WHILE N !! b + INV { invar2 (list A) (list B) (list N) ai a \ b = match_array A a \ z = n-ai } VAR {var (list A) { stable (list A) { opti\<^sub>a (list A)]" proof (vcg_tc, goal_cases) case 1 thus ?case (* outer invar holds initially *) by(auto simp: pref_match_def P_set card_distinct match_def list_array index_nth_id prefers_def opti\<^sub>a_def \_def cong: conj_cong) next case 2 (* outer invar and b implies inner invar *) thus ?case by (auto simp: atLeastLessThanSuc_atLeastAtMost simp flip: atLeastLessThan_eq_atLeastAtMost_diff) next case 3 (* preservation of inner invar *) have R: "\ba1a2 Q ! b \ a1 < a2" by (simp add: Q_set \R = _\ length_Q ranking_iff_pref) show ?case proof (simp only: mem_Collect_eq prod.case, goal_cases) case 1 show ?case using inner_pres[OF R _ _ refl refl refl] 3 unfolding array_abs by blast qed next case 4 (* inner invar' and not b implies outer invar *) show ?case proof (simp only: mem_Collect_eq prod.case, rule conjI, goal_cases) case 1 show ?case using 4 inner_to_outer unfolding array_abs by blast next case 2 thus ?case using 4 by auto qed next case 5 (* outer invar and not b ibplies post *) thus ?case using pref_match_stable unfolding invAM_def invar1_def by(metis le_neq_implies_less) qed subsection \Algorithm 7.1: single-loop variant\ lemma Gale_Shapley7_1: assumes "R = map ranking Q" -shows "VARS A B M a a' ai b r - [ai = 0 \ a = 0 \ A = array 0 n \ B = array 0 n \ M = array False n] +shows "VARS A B N a a' ai b r + [ai = 0 \ a = 0 \ A = array 0 n \ B = array 0 n \ N = array False n] WHILE ai < n - INV { invar2' (list A) (list B) (list M) ai a } + INV { invar2' (list A) (list B) (list N) ai a } VAR {var (list A) ({ M !! b - THEN B := B[b ::= a]; M := M[b ::= True]; ai := ai + 1; a := ai + IF \ N !! b + THEN B := B[b ::= a]; N := N[b ::= True]; ai := ai + 1; a := ai ELSE a' := B !! b; r := R ! match_array A a'; IF r ! a < r ! a' THEN B := B[b ::= a]; A := A[a' ::= A!!a' + 1]; a := a' ELSE A := A[a ::= A!!a + 1] FI FI OD [matching (list A) { stable (list A) { opti\<^sub>a (list A)]" proof (vcg_tc, goal_cases) case 1 thus ?case by(auto simp: pref_match_def P_set card_distinct match_def list_array index_nth_id prefers_def opti\<^sub>a_def \_def cong: conj_cong) next case 3 thus ?case using pref_match_stable atLeast0_lessThan_Suc[of n] by force next - case (2 v A B M a a' ai) - have R': "M !! match_array A a \ + case (2 v A B N a a' ai) + have R': "N !! match_array A a \ (R ! match_array A (B !! match_array A a) ! a < R ! match_array A (B !! match_array A a) ! (B !! match_array A a)) = (Q ! match_array A (B !! match_array A a) \ a < B !! match_array A a)" using R_iff_P 2 assms by (metis array_abs) show ?case apply(simp only:mem_Collect_eq prod.case) - using 2 R' pres2'[of "list A" "list B" "list M" ai a] by (metis array_abs) + using 2 R' pres2'[of "list A" "list B" "list N" ai a] by (metis array_abs) qed end subsection \Executable functional Code\ definition gs_inner where -"gs_inner P R M = - while (\(A,B,a,b). M !! b) +"gs_inner P R N = + while (\(A,B,a,b). N !! b) (\(A,B,a,b). let a' = B !! b; r = R !! (P !! a' !! (A !! a')) in let (A, B, a) = if r !! a < r !! a' then (A[a' ::= A !! a' + 1], B[b ::= a], a') else (A[a ::= A !! a + 1], B, a) in (A, B, a, P !! a !! (A !! a)))" definition gs :: "nat \ nat array array \ nat array array \ nat array \ nat array \ bool array \ nat" where "gs n P R = - while (\(A,B,M,ai). ai < n) - (\(A,B,M,ai). - let (A,B,a,b) = gs_inner P R M (A, B, ai, P !! ai !! (A !! ai)) - in (A, B[b ::= a], M[b::=True], ai+1)) + while (\(A,B,N,ai). ai < n) + (\(A,B,N,ai). + let (A,B,a,b) = gs_inner P R N (A, B, ai, P !! ai !! (A !! ai)) + in (A, B[b ::= a], N[b::=True], ai+1)) (array 0 n, array 0 n, array False n, 0)" definition gs1 :: "nat \ nat array array \ nat array array \ nat array \ nat array \ bool array \ nat \ nat" where "gs1 n P R = - while (\(A,B,M,ai,a). ai < n) - (\(A,B,M,ai,a). + while (\(A,B,N,ai,a). ai < n) + (\(A,B,N,ai,a). let b = P !! a !! (A !! a) - in if \ M !! b - then (A, B[b ::= a], M[b ::= True], ai+1, ai+1) + in if \ N !! b + then (A, B[b ::= a], N[b ::= True], ai+1, ai+1) else let a' = B !! b; r = R !! (P !! a' !! (A !! a')) in if r !! a < r !! a' - then (A[a' ::= A!!a' + 1], B[b ::= a], M, ai, a') - else (A[a ::= A!!a + 1], B, M, ai, a)) + then (A[a' ::= A!!a' + 1], B[b ::= a], N, ai, a') + else (A[a ::= A!!a + 1], B, N, ai, a)) (array 0 n, array 0 n, array False n, 0, 0)" definition "pref_array = array_of_list o map array_of_list" lemma list_list_pref_array: "i < length Pa \ list (list (pref_array Pa) ! i) = Pa ! i" by(simp add: pref_array_def) fun rk_of_pref :: "nat \ nat array \ nat list \ nat array" where "rk_of_pref r rs (n#ns) = (rk_of_pref (r+1) rs ns)[n ::= r]" | "rk_of_pref r rs [] = rs" definition rank_array1 :: "nat list \ nat array" where "rank_array1 P = rk_of_pref 0 (array 0 (length P)) P" definition rank_array :: "nat list list \ nat array array" where "rank_array = array_of_list o map rank_array1" lemma length_rk_of_pref[simp]: "array_length(rk_of_pref v vs P) = array_length vs" by(induction P arbitrary: v)(auto) lemma nth_rk_of_pref: "\ length P \ array_length rs; i \ set P; distinct P; set P \ { \ rk_of_pref r rs P !! i = index P i + r" by(induction P arbitrary: r i) (auto simp add: array_get_array_set_other) lemma rank_array1_iff_pref: "\ set P = { \ rank_array1 P !! i < rank_array1 P !! j \ P \ i < j" by(simp add: rank_array1_def prefers_def nth_rk_of_pref card_distinct) definition Gale_Shapley where "Gale_Shapley P Q = (if Pref P Q then Some (fst (gs (length P) (pref_array P) (rank_array Q))) else None)" definition Gale_Shapley1 where "Gale_Shapley1 P Q = (if Pref P Q then Some (fst (gs1 (length P) (pref_array P) (rank_array Q))) else None)" (*export_code Gale_Shapley_array in SML*) context Pref begin lemma gs_inner: assumes "R = rank_array Q" -assumes "invar2 (list A) (list B) (list M) ai a" "b = match_array A a" -shows "gs_inner (pref_array P) R M (A, B, a, b) = (A',B',a',b') - \ invar1 (list A') ((list B')[b' := a']) ((list M)[b' := True]) (ai+1)" +assumes "invar2 (list A) (list B) (list N) ai a" "b = match_array A a" +shows "gs_inner (pref_array P) R N (A, B, a, b) = (A',B',a',b') + \ invar1 (list A') ((list B')[b' := a']) ((list N)[b' := True]) (ai+1)" unfolding gs_inner_def proof(rule while_rule2[where - P = "\(A,B,a,b). invar2 (list A) (list B) (list M) ai a \ b = match_array A a" + P = "\(A,B,a,b). invar2 (list A) (list B) (list N) ai a \ b = match_array A a" and r = "measure (\(A, B, a, b). Pref.var0 P (list A) {ba1a2 Q ! b \ a1 < a2" using rank_array1_iff_pref by(simp add: \R = _\ length_Q array_get Q_set rank_array_def) have ***: "match (list A) (list B ! b) < length (list R)" using s inv(1)[unfolded invar2_def] using ** by(simp add: \R = _\ rank_array_def match_array match_less_n length_Q) show ?case using inv apply(simp only: s prod.case Let_def split: if_split) - using inner_pres[OF R _ _ refl refl refl refl refl, of "list A" "list B" "list M" ai a b] + using inner_pres[OF R _ _ refl refl refl refl refl, of "list A" "list B" "list N" ai a b] unfolding invar2_def array_abs list_list_pref_array[OF **[unfolded n_def]] list_list_pref_array[OF *[unfolded n_def]] nth_map[OF ***] unfolding match_def by presburger case 2 show ?case using inv apply(simp only: s prod.case Let_def in_measure split: if_split) - using inner_pres2[OF R _ _ refl refl refl refl refl, of "list A" "list B" "list M" ai a b] + using inner_pres2[OF R _ _ refl refl refl refl refl, of "list A" "list B" "list N" ai a b] unfolding invar2_def array_abs list_list_pref_array[OF **[unfolded n_def]] list_list_pref_array[OF *[unfolded n_def]] nth_map[OF ***] unfolding match_def by presburger qed next case 3 show ?case proof (rule, goal_cases) case 1 show ?case by(rule inner_to_outer[OF 3[unfolded 1 prod.case array_abs]]) qed next case 4 show ?case by simp qed lemma gs: assumes "R = rank_array Q" -shows "gs n (pref_array P) R = (A,B,M,ai) \ matching (list A) { stable (list A) { opti\<^sub>a (list A)" +shows "gs n (pref_array P) R = (A,B,N,ai) \ matching (list A) { stable (list A) { opti\<^sub>a (list A)" unfolding gs_def -proof(rule while_rule2[where P = "\(A,B,M,ai). invar1 (list A) (list B) (list M) ai" - and r = "measure(\(A,B,M,ai). n - ai)"], goal_cases) +proof(rule while_rule2[where P = "\(A,B,N,ai). invar1 (list A) (list B) (list N) ai" + and r = "measure(\(A,B,N,ai). n - ai)"], goal_cases) case 1 show ?case by(auto simp: pref_match_def P_set card_distinct match_def list_array index_nth_id prefers_def opti\<^sub>a_def \_def cong: conj_cong) next case (2 s) - obtain A B M ai where s: "s = (A, B, M, ai)" + obtain A B N ai where s: "s = (A, B, N, ai)" using prod_cases4 by blast - have 1: "invar2 (list A) (list B) (list M) ai ai" using 2 s + have 1: "invar2 (list A) (list B) (list N) ai ai" using 2 s by (auto simp: atLeastLessThanSuc_atLeastAtMost simp flip: atLeastLessThan_eq_atLeastAtMost_diff) hence "ai < n" by(simp) show ?case using 2 s gs_inner[OF \R = _ \ 1] by (auto simp: array_abs match_def list_list_pref_array[OF \ai < n\[unfolded n_def]] simp del: invar1_def split: prod.split) next case 3 thus ?case using pref_match_stable by auto next case 4 show ?case by simp qed lemma R_iff_P: -assumes "R = rank_array Q" "invar2' A B M ai a" "ai < n" "M ! match A a" +assumes "R = rank_array Q" "invar2' A B N ai a" "ai < n" "N ! match A a" "b = match A a" "a' = B ! b" shows "(list (list R ! match A a') ! a < list (list R ! match A a') ! a') = (Q ! match A a' \ a < a')" proof - have R: "\ba1a2 Q ! b \ a1 < a2" by (simp add: Q_set \R = _\ length_Q array_of_list_def rank_array_def rank_array1_iff_pref) let ?M = "{ { { ?M" - using invAB match_less_n[OF A] as \M!match A a\ by (metis \_Some ranI) + using invAB match_less_n[OF A] as \N!match A a\ by (metis \_Some ranI) hence "B ! match A a < n" using M by auto thus ?thesis using assms match_less_n R by simp (metis array_get as) qed lemma gs1: assumes "R = rank_array Q" -shows "gs1 n (pref_array P) R = (A,B,M,ai,a) \ matching (list A) { stable (list A) { opti\<^sub>a (list A)" +shows "gs1 n (pref_array P) R = (A,B,N,ai,a) \ matching (list A) { stable (list A) { opti\<^sub>a (list A)" unfolding gs1_def -proof(rule while_rule2[where P = "\(A,B,M,ai,a). invar2' (list A) (list B) (list M) ai a" - and r = "measure(\(A,B,M,ai,a). Pref.var P (list A) ({a_def \_def cong: conj_cong) next case (2 s) - obtain A B M ai a where s: "s = (A, B, M, ai, a)" + obtain A B N ai a where s: "s = (A, B, N, ai, a)" using prod_cases5 by blast - have 1: "invar2' (list A) (list B) (list M) ai a" using 2(1) s + have 1: "invar2' (list A) (list B) (list N) ai a" using 2(1) s by (auto simp: atLeastLessThanSuc_atLeastAtMost simp flip: atLeastLessThan_eq_atLeastAtMost_diff) have "ai < n" using 2(2) s by(simp) hence "a < n" using 1 by simp hence "match (list A) a < n" using 1 match_less_n by auto - hence *: "list M ! match (list A) a \ list B ! match (list A) a < n" + hence *: "list N ! match (list A) a \ list B ! match (list A) a < n" using s 1[unfolded invar2'_def] apply (simp add: array_abs ran_def) using atLeast0LessThan by blast - have R': "list M ! match (list A) a \ + have R': "list N ! match (list A) a \ (list (list R ! match (list A) (list B ! match (list A) a)) ! a < list (list R ! match (list A) (list B ! match (list A) a)) ! (list B ! match (list A) a)) = (Q ! match (list A) (list B ! match (list A) a) \ a < list B ! match (list A) a)" using R_iff_P \R = _\ 1 \ai < n\ by blast show ?case using s apply (simp add: Let_def) unfolding list_list_pref_array[OF \a < n\[unfolded n_def]] array_abs using list_list_pref_array[OF *[unfolded n_def]] (* conditional, cannot unfold :-( *) pres2'[OF 1 \ai < n\ refl refl refl refl refl] R' apply(intro conjI impI) by (auto simp: match_def) (* w/o intro too slow because of conditional eqn above *) next case 3 thus ?case using pref_match_stable atLeast0_lessThan_Suc[of n] by force next case 4 show ?case by simp qed end theorem gs: "\ Pref P Q; n = length P \ \ \A. Gale_Shapley P Q = Some A \ Pref.matching P (list A) { Pref.stable P Q (list A) { Pref.opti\<^sub>a P Q (list A)" unfolding Gale_Shapley_def using Pref.gs by (metis fst_conv surj_pair) theorem gs1: "\ Pref P Q; n = length P \ \ \A. Gale_Shapley1 P Q = Some A \ Pref.matching P (list A) { Pref.stable P Q (list A) { Pref.opti\<^sub>a P Q (list A)" unfolding Gale_Shapley1_def using Pref.gs1 by (metis fst_conv surj_pair) text \Two examples from Gusfield and Irving:\ lemma "list_of_array (the (Gale_Shapley [[3,0,1,2], [1,2,0,3], [1,3,2,0], [2,0,3,1]] [[3,0,2,1], [0,2,1,3], [0,1,2,3], [3,0,2,1]])) = [0,1,0,1]" by eval lemma "list_of_array (the (Gale_Shapley [[4,6,0,1,5,7,3,2], [1,2,6,4,3,0,7,5], [7,4,0,3,5,1,2,6], [2,1,6,3,0,5,7,4], [6,1,4,0,2,5,7,3], [0,5,6,4,7,3,1,2], [1,4,6,5,2,3,7,0], [2,7,3,4,6,1,5,0]] [[4,2,6,5,0,1,7,3], [7,5,2,4,6,1,0,3], [0,4,5,1,3,7,6,2], [7,6,2,1,3,0,4,5], [5,3,6,2,7,0,1,4], [1,7,4,2,3,5,6,0], [6,4,1,0,7,5,3,2], [6,3,0,4,1,2,5,7]])) = [0, 1, 0, 5, 0, 0, 0, 2]" by eval end \ No newline at end of file diff --git a/thys/Solidity/Reentrancy.thy b/thys/Solidity/Reentrancy.thy --- a/thys/Solidity/Reentrancy.thy +++ b/thys/Solidity/Reentrancy.thy @@ -1,840 +1,838 @@ 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" +lemma safe_external: "Qe ad iv (st::State)" 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" +lemma safe_fallback: "Qfe ad iv st" 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) + assumes "\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)))" - apply (rule sound) using assms safe_external safe_fallback safe_constructor by auto + apply (rule invariant) using assms safe_external safe_fallback safe_constructor by auto end end diff --git a/thys/Solidity/Weakest_Precondition.thy b/thys/Solidity/Weakest_Precondition.thy --- a/thys/Solidity/Weakest_Precondition.thy +++ b/thys/Solidity/Weakest_Precondition.thy @@ -1,2200 +1,2172 @@ 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 + 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 simp 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 + 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'\ + assumes "type (accounts st ad) = Some (Contract cname)" + and "\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'\ + assumes "type (accounts st ad) = Some (Contract cname)" + and "\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'\ + assumes "type (accounts st ad) = Some (Contract cname)" + and "\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'\ + assumes "type (accounts st ad) = Some (Contract cname)" + and "\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: +lemma invariant_rec: fixes iv ad - assumes "\ad. external ad iv" - and "\ad. fallback ad iv" + assumes "\ad (st::State). Qe ad iv st" + and "\ad (st::State). Qfe ad iv st" 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 + moreover from `type (acc ad) = Some (Contract c)` 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 "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)] by linarith 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 - + from `c = cname` have "type (accounts st ad) = Some (Contract cname)" using 1(4) True by simp 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 + (st\gas := g'', accounts := acc, stack := k\<^sub>l, memory := m\<^sub>l\)" using 1 True using assms(1) 1(8) 7(3) unfolding Qe_def by simp 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 + moreover from `type (acc ad) = Some (Contract c)` 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 + from `c = cname` have "type (accounts st ad) = Some (Contract cname)" using 2(4) True by simp + then have "iv (storage st ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st ad)))" using 7(4) 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 + moreover have "g' \ gas st" using msel_ssel_expr_load_rexp_gas(3)[OF 2(2)] msel_ssel_expr_load_rexp_gas(3)[OF 2(6)] by linarith 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 + (st\gas := g', accounts := acc, stack := emptyStore, memory := emptyStore\)" using assms(2) 7(3) 2 True unfolding Qfe_def by simp 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 + moreover from `type (acc ad) = Some (Contract c)` 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 "g' \ gas st" using msel_ssel_expr_load_rexp_gas(3)[OF 1(2)] msel_ssel_expr_load_rexp_gas(3)[OF 1(3)] by linarith 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 + from `c = cname` have "type (accounts st ad) = Some (Contract cname)" using 1(5) True by simp + then have "iv (storage st ad) (ReadL\<^sub>i\<^sub>n\<^sub>t (bal (accounts st ad)))" using 8(3) 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 + (st\gas := g', accounts := acc, stack := emptyStore, memory := emptyStore\)" using assms(2) 8(2) 1 True unfolding Qfe_def by simp 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: +theorem invariant: fixes iv ad - assumes "\ad. external ad iv" - and "\ad. fallback ad iv" + assumes "\ad (st::State). Qe ad iv st" + and "\ad (st::State). Qfe ad iv st" 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 + using assms invariant_rec 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, + unfold Qe_def, elims, (erule cases;simp) method fallback uses cases = - unfold fallback_def, + unfold Qfe_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], + solve simp, (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