Theory Storage
section‹Storage›
theory Storage
imports Valuetypes "Finite-Map-Extras.Finite_Map_Extras"
begin
subsection ‹Hashing›
definition hash :: "Location ⇒ String.literal ⇒ Location"
  where "hash loc ix = ix + (STR ''.'' + loc)"
declare hash_def [solidity_symbex]
lemma example: "hash (STR ''1.0'') (STR ''2'') = hash (STR ''0'') (STR ''2.1'')" by eval
lemma hash_explode:
  "String.explode (hash l i) = String.explode i @ (String.explode (STR ''.'') @ String.explode l)"
  unfolding hash_def by (simp add: plus_literal.rep_eq)
lemma hash_dot:
  "String.explode (hash l i) ! length (String.explode i) = CHR ''.''"
  unfolding hash_def by (simp add: Literal.rep_eq plus_literal.rep_eq)
lemma hash_injective:
  assumes "hash l i = hash l' i'"
    and "CHR ''.'' ∉ set (String.explode i)"
    and "CHR ''.'' ∉ set (String.explode i')"
  shows "l = l' ∧ i = i'"
proof (rule ccontr)
  assume "¬ (l = l' ∧ i = i')" 
  then consider (1) "i≠i'" | (2) "i=i' ∧ l≠l'" by auto
  then have "String.explode (hash l i) ≠ String.explode (hash l' i')"
  proof cases
    case 1
    then have neq: "(String.explode i) ≠ (String.explode i')" by (metis literal.explode_inverse)
    consider (1) "length (String.explode i) = length (String.explode i')" | (2) "length (String.explode i) < length (String.explode i')" | (3) "length (String.explode i) > length (String.explode i')" by linarith
    then show ?thesis
    proof (cases)
      case 1
      then obtain j where "String.explode i ! j ≠ String.explode i' ! j" using neq nth_equalityI by auto
      then show ?thesis using 1 plus_literal.rep_eq unfolding hash_def by force
    next
      case 2
      then have "String.explode i' ! length (String.explode i) ≠ CHR ''.''" using assms(3) by (metis nth_mem)
      then have "String.explode (hash l' i') ! length (String.explode i) ≠ CHR ''.''" using 2 hash_explode[of l' i'] by (simp add: nth_append)
      moreover have "String.explode (hash l i) ! length (String.explode i) = CHR ''.''" using hash_dot by simp
      ultimately show ?thesis by auto
    next
      case 3
      then have "String.explode i ! length (String.explode i') ≠ CHR ''.''" using assms(2) by (metis nth_mem)
      then have "String.explode (hash l i) ! length (String.explode i') ≠ CHR ''.''" using 3 hash_explode[of l i] by (simp add: nth_append)
      moreover have "String.explode (hash l' i') ! length (String.explode i') = CHR ''.''" using hash_dot by simp
      ultimately show ?thesis by auto
    qed
  next
    case 2
    then show ?thesis using hash_explode literal.explode_inject by force
  qed
  then show False using assms(1) by simp
qed
subsection ‹General Store›
record 'v Store =
  mapping :: "(Location,'v) fmap"
  toploc :: nat 
definition accessStore :: "Location ⇒ 'v Store ⇒ 'v option"
where "accessStore loc st = fmlookup (mapping st) loc"
declare accessStore_def[solidity_symbex]
definition emptyStore :: "'v Store"
where "emptyStore = ⦇ mapping=fmempty, toploc=0 ⦈"
declare emptyStore_def [solidity_symbex]
definition allocate :: "'v Store ⇒ Location * ('v Store)"
where "allocate s = (let ntop = Suc(toploc s) in (ShowL⇩n⇩a⇩t ntop, s ⦇toploc := ntop⦈))"
definition updateStore :: "Location ⇒ 'v ⇒ 'v Store ⇒ 'v Store"
where "updateStore loc val s = s ⦇ mapping := fmupd loc val (mapping s)⦈"
declare updateStore_def [solidity_symbex]
definition push :: "'v ⇒ 'v Store ⇒ 'v Store"
  where "push val sto = (let s = updateStore (ShowL⇩n⇩a⇩t (toploc sto)) val sto in snd (allocate s))"
declare push_def [solidity_symbex]
subsection ‹Stack›