Theory Fun4
section "Proof of Relative Security for fun4"
theory Fun4
imports "../Instance_IMP/Instance_Secret_IMem"
"Relative_Security.Unwinding_fin"
begin
subsection "Function definition and Boilerplate"
no_notation bot (‹⊥›)
consts NN :: nat
consts size_aa1 :: nat
consts size_aa2 :: nat
lemma NN: "int NN ≥ 0" by auto
locale array_nempty = assumes aa1:"size_aa1 > 0" and NN: "int NN > 0"
definition aa1 :: "avname" where "aa1 = ''a1''"
definition aa2 :: "avname" where "aa2 = ''a2''"
definition vv :: "avname" where "vv = ''v''"
definition xx :: "avname" where "xx = ''i''"
definition tt :: "avname" where "tt = ''w''"
lemmas vvars_defs = aa1_def aa2_def vv_def xx_def tt_def
lemma vvars_dff[simp]:
"aa1 ≠ aa2" "aa1 ≠ vv" "aa1 ≠ xx" "aa1 ≠ tt"
"aa2 ≠ aa1" "aa2 ≠ vv" "aa2 ≠ xx" "aa1 ≠ tt"
"vv ≠ aa1" "vv ≠ aa2" "vv ≠ xx" "vv ≠ tt"
"xx ≠ aa1" "xx ≠ aa2" "xx ≠ vv" "xx ≠ tt"
"tt ≠ aa1" "tt ≠ aa2" "tt ≠ vv" "tt ≠ xx"
unfolding vvars_defs by auto
fun initAvstore :: "avstore ⇒ bool" where
"initAvstore (Avstore as) = (as aa1 = (0, size_aa1) ∧ as aa2 = (size_aa1, size_aa2))"
fun istate ::"state ⇒ bool" where
"istate s = (initAvstore (getAvstore s))"
definition "prog ≡
[
Start ,
Input U xx ,
tt ::= (N 0) ,
IfJump (Less (V xx) (N NN)) 4 6 ,
vv ::= VA aa1 (N 0) ,
tt ::= Plus (VA aa2 (Times (V vv) (N 512))) (V xx) ,
Output U (V tt)
]"
lemma cases_6: "(i::pcounter) = 0 ∨ i = 1 ∨ i = 2 ∨ i = 3 ∨ i = 4 ∨ i = 5 ∨
i = 6 ∨ i > 6"
apply(cases i, simp_all)
subgoal for i apply(cases i, simp_all)
subgoal for i apply(cases i, simp_all)
subgoal for i apply(cases i, simp_all)
subgoal for i apply(cases i, simp_all)
subgoal for i apply(cases i, simp_all)
subgoal for i apply(cases i, simp_all)
. . . . . . .
lemma cases_thenBranch: "(i::pcounter) < 4 ∨ i = 4 ∨ i = 5 ∨ i = 6 ∨ i > 6"
apply(cases i, simp_all)
subgoal for i apply(cases i, simp_all)
subgoal for i apply(cases i, simp_all)
subgoal for i apply(cases i, simp_all)
subgoal for i apply(cases i, simp_all)
subgoal for i apply(cases i, simp_all)
subgoal for i apply(cases i, simp_all)
. . . . . . .
lemma xx_NN_cases: "vs xx < int NN ∨ vs xx ≥ int NN" by auto
lemma is_If_pcOf[simp]:
"pcOf cfg < 7 ⟹ is_IfJump (prog ! (pcOf cfg)) ⟷ pcOf cfg = 3"
apply(cases cfg) subgoal for pc s using cases_6[of "pcOf cfg"]
by (auto simp: prog_def) .
lemma is_If_pc[simp]:
"pc < 7 ⟹ is_IfJump (prog ! pc) ⟷ pc = 3"
using cases_6[of pc]
by (auto simp: prog_def)
lemma is_If_pcThen[simp]: "pcOf cfg ∈ {4..6} ⟹ ¬is_IfJump (prog ! pcOf cfg)"
using cases_thenBranch[of "pcOf cfg"]
by (auto simp: prog_def)
consts mispred :: "predState ⇒ pcounter list ⇒ bool"
fun resolve :: "predState ⇒ pcounter list ⇒ bool" where
"resolve p pc = (if (pc = [6,6] ∨ pc = [4,6]) then True else False)"
consts update :: "predState ⇒ pcounter list ⇒ predState"
consts initPstate :: predState