Theory Fun2

section "Proof of Relative Security for fun2"
theory Fun2
  imports  
  "../Instance_IMP/Instance_Secret_IMem"
  "Relative_Security.Unwinding_fin" 
begin 

subsection "Function definition and Boilerplate"
no_notation bot ("")

consts NN :: nat
lemma NN: "NN  0" by auto


definition aa1 :: "avname" where "aa1 = ''a1''" 
definition aa2 :: "avname" where "aa2 = ''a2''" 
definition xx :: "avname" where "xx = ''xx''" 
definition tt :: "avname" where "tt = ''tt''" 

lemmas vvars_defs = aa1_def aa2_def xx_def tt_def

lemma vvars_dff[simp]: 
"aa1  aa2" "aa1  xx" "aa1  tt"
"aa2  aa1" "aa2  xx" "aa2  tt"
"xx  aa1" "xx  aa2" "xx  tt"
"tt  aa1" "tt  aa2" "tt  xx"
unfolding vvars_defs by auto

consts size_aa1 :: nat
consts size_aa2 :: nat

lemma aa1: "size_aa1  0" and aa2:"size_aa2  0" by auto

(* The initial avstore contains two arrays named aa1 and aa2 located one after the other *)
fun initAvstore :: "avstore  bool" where
"initAvstore (Avstore as)  = (as aa1 = (0, nat size_aa1)  as aa2 = (nat size_aa1, nat size_aa2))"

fun istate ::"state  bool" where
"istate s = (initAvstore (getAvstore s))"


definition "prog  
[
 ⌦‹0 › Start , 
 ⌦‹1 › Input U xx ,    
 ⌦‹2 › tt ::= (N 0) ,
 ⌦‹3 › IfJump (Less (V xx) (N NN)) 4 6 ,
 ⌦‹4 ›   Fence ,
 ⌦‹5 ›   tt ::= (VA aa2 (Times (VA aa1 (V xx)) (N 512))),
 ⌦‹6 › 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 xx_NN_cases: "vs xx < int(NN)  vs xx  int(NN)" by auto

lemma is_If_pcOf[simp]: 
"pcOf cfg < 6  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 < 6  is_IfJump (prog ! pc)  pc = 3"
using cases_6[of pc] 
by (auto simp: prog_def) 

lemma eq_Fence_pc[simp]: 
"pc < 6  prog ! pc = Fence  pc = 4"
using cases_6[of pc] 
by (auto simp: prog_def) 


(* *)

consts mispred :: "predState  pcounter list  bool"
fun resolve :: "predState  pcounter list  bool" where
  "resolve p pc = (if (set pc = {6,4}) then True else False)"


consts update :: "predState  pcounter list  predState"
consts initPstate :: predState

(* *)
interpretation Prog_Mispred_Init where 
prog = prog and initPstate = initPstate and 
mispred = mispred and resolve = resolve and update = update and 
istate = istate
  by (standard, simp add: prog_def)

(* *)

(* Restoring the abbreviations: *)
abbreviation
  stepB_abbrev :: "config × val llist × val llist  config × val llist × val llist  bool" (infix "→B" 55)
  where "x →B y == stepB x y"

abbreviation
  stepsB_abbrev :: "config × val llist × val llist  config × val llist × val llist  bool" (infix "→B*" 55)
  where "x →B* y == star stepB x y"

abbreviation
  stepM_abbrev :: "config × val llist × val llist  config × val llist × val llist  bool" (infix "→M" 55)
  where "x →M y == stepM x y"

abbreviation
  stepN_abbrev :: "config × val llist × val llist × loc set  config × val llist × val llist × loc set  bool" (infix "→N" 55)
  where "x →N y == stepN x y"

abbreviation
  stepsN_abbrev :: "config × val llist × val llist × loc set  config × val llist × val llist × loc set  bool" (infix "→N*" 55)
  where "x →N* y == star stepN x y"

abbreviation
  stepS_abbrev :: "configS  configS  bool" (infix "→S" 55)
  where "x →S y == stepS x y"

abbreviation
  stepsS_abbrev :: "configS  configS  bool" (infix "→S*" 55)
  where "x →S* y == star stepS x y"

(* *)

lemma endPC[simp]: "endPC = 7"
unfolding endPC_def unfolding prog_def by auto

(* *)

lemma is_getUntrustedInput_pcOf[simp]: "pcOf cfg < 6  is_getInput (prog!(pcOf cfg))  pcOf cfg = 1"
  using cases_6[of "pcOf cfg"] by (auto simp: prog_def)

lemma start[simp]:"prog ! 0 = Start"
  by (auto simp: prog_def)

lemma getUntrustedInput_pcOf[simp]: " prog!1 = Input U xx"
  by (auto simp: prog_def)


lemma if_stat[simp]: "prog! 3 = (IfJump (Less (V xx) (N NN)) 4 6)"
  by (auto simp: prog_def)

lemma isOutput1[simp]:"prog ! 6 = Output U (V tt)"
  by (auto simp: prog_def)


lemma is_Output_pcOf[simp]: "pcOf cfg < 6  is_Output (prog!(pcOf cfg))  pcOf cfg = 6"
using cases_6[of "pcOf cfg"] by (auto simp: prog_def)


lemma is_Fence_pcOf[simp]: "pcOf cfg < 6 (prog!(pcOf cfg)) = Fence  pcOf cfg = 4"
using cases_6[of "pcOf cfg"] by (auto simp: prog_def)

lemma is_Output[simp]: "is_Output (prog ! 6)"
unfolding is_Output_def prog_def by auto


(* *)

lemma isSecV_pcOf[simp]: 
"isSecV (cfg,ibT, ibUT)  pcOf cfg = 0"
using isSecV_def by simp

lemma isSecO_pcOf[simp]: 
"isSecO (pstate,cfg,cfgs,ibT, ibUT,ls)  (pcOf cfg = 0  cfgs = [])"
using isSecO_def by simp

(* *)

lemma getInputT_not[simp]: "pcOf cfg < 7  
(prog ! pcOf cfg)  Input T inp"
apply(cases cfg) subgoal for pc s using cases_6[of "pcOf cfg "] 
by (auto simp: prog_def) .

lemma getActV_pcOf[simp]: 
"pcOf cfg < 7  
 getActV (cfg,ibT,ibUT,ls) = 
 (if pcOf cfg = 1 then lhd ibUT else )"
apply(subst getActV_simps) unfolding prog_def 
using cases_6[of "pcOf cfg"] by auto

lemma getObsV_pcOf[simp]: 
"pcOf cfg < 7   
 getObsV (cfg,ibT,ibUT,ls) = 
 (if pcOf cfg = 6 then 
  (outOf (prog!(pcOf cfg)) (stateOf cfg), ls)
  else  
 )"
apply(subst getObsV_simps) 
using getObsV_simps not_is_Output_getObsV is_Output_pcOf 
unfolding prog_def by simp  


lemma getActO_pcOf[simp]: 
"pcOf cfg < 7  
 getActO (pstate,cfg,cfgs,ibT,ibUT,ls) = 
 (if pcOf cfg = 1  cfgs = [] then lhd ibUT else )"
apply(subst getActO_simps)
apply(cases cfgs, auto)
unfolding prog_def apply simp 
using getActV_simps getActV_pcOf prog_def by presburger

lemma getObsO_pcOf[simp]: 
"pcOf cfg < 7   
 getObsO (pstate,cfg,cfgs,ibT, ibUT,ls) = 
 (if (pcOf cfg = 6  cfgs = []) then 
  (outOf (prog!(pcOf cfg)) (stateOf cfg), ls)
  else  
 )"
apply(subst getObsO_simps) 
  apply(cases cfgs, auto)
using getObsV_simps is_Output_pcOf not_is_Output_getObsV
unfolding prog_def by auto

(* *)


lemma eqSec_pcOf[simp]: 
"eqSec (cfg1, ibT, ibUT1, ls1) (pstate3, cfg3, cfgs3, ibT, ibUT3, ls3)  
 (pcOf cfg1 = 0  pcOf cfg3 = 0  cfgs3 = [])  
 (pcOf cfg1 = 0  stateOf cfg1 = stateOf cfg3)"
unfolding eqSec_def by simp


(* nextB, nextM and readLocs behavior 
(for nextM we are only interested at branching points, here only program counter 4): *)

lemma nextB_pc0[simp]: 
"nextB (Config 0 s, ibT, ibUT) = 
 (Config 1 s, ibT, ibUT)"
apply(subst nextB_Start_Skip_Fence)
  unfolding endPC_def unfolding prog_def by auto

lemma nextB_pc0'[simp]:"nextB (Config 0 (State (Vstore vs) avst h p), ibT, ibUT) = 
       (Config (Suc 0) (State (Vstore vs) avst h p), ibT, ibUT)"
apply(subst nextB_Start_Skip_Fence)
  unfolding endPC_def unfolding prog_def by auto



lemma readLocs_pc0[simp]: 
"readLocs (Config 0 s) = {}"
unfolding endPC_def readLocs_def unfolding prog_def by auto

(* *)

lemma nextB_pc1[simp]: 
"ibUT  LNil  nextB (Config 1 (State (Vstore vs) avst h p), ibT, ibUT) = 
 (Config 2 (State (Vstore (vs(xx := lhd ibUT))) avst h p), ibT, ltl ibUT)"
apply(subst nextB_getUntrustedInput')
unfolding endPC_def unfolding prog_def by auto

lemma readLocs_pc1[simp]: 
"readLocs (Config 1 s) = {}"
unfolding endPC_def readLocs_def unfolding prog_def by auto

lemma nextB_pc1'[simp]: 
"ibUT  LNil  nextB (Config (Suc 0) (State (Vstore vs) avst h p), ibT, ibUT) = 
 (Config 2 (State (Vstore (vs(xx := lhd ibUT))) avst h p), ibT, ltl ibUT)"
apply(subst nextB_getUntrustedInput')
unfolding endPC_def unfolding prog_def by auto

lemma readLocs_pc1'[simp]: 
"readLocs (Config (Suc 0) s) = {}"
unfolding endPC_def readLocs_def unfolding prog_def by auto

(* *)

lemma nextB_pc2[simp]:
"nextB (Config 2 (State (Vstore vs) avst h p), ibT, ibUT) = 
 (Config 3 (State (Vstore (vs(tt := 0))) avst h p), ibT, ibUT)"
apply(subst nextB_Assign)
unfolding endPC_def unfolding prog_def by auto

lemma readLocs_pc2[simp]: 
"readLocs (Config 2 s) = {}"
unfolding endPC_def readLocs_def unfolding prog_def by auto

(* *)

lemma nextB_pc3_then[simp]: 
"vs xx < NN 
 nextB (Config 3 (State (Vstore vs) avst h p), ibT, ibUT) = 
 (Config 4 (State (Vstore vs) avst h p), ibT, ibUT)"
apply(subst nextB_IfTrue)
unfolding endPC_def unfolding prog_def by auto

lemma nextB_pc3_else[simp]: 
"vs xx  NN 
 nextB (Config 3 (State (Vstore vs) avst h p), ibT, ibUT) = 
 (Config 6 (State (Vstore vs) avst h p), ibT, ibUT)"
apply(subst nextB_IfFalse)
unfolding endPC_def unfolding prog_def by auto

lemma nextB_pc3: 
"nextB (Config 3 (State (Vstore vs) avst h p), ibT, ibUT) = 
 (Config (if vs xx < NN then 4 else 6) (State (Vstore vs) avst h p), ibT, ibUT)"
  by(cases "vs xx < NN", auto)

lemma nextM_pc3_then[simp]: 
"vs xx  NN 
 nextM (Config 3 (State (Vstore vs) avst h p), ibT, ibUT) = 
 (Config 4 (State (Vstore vs) avst h p), ibT, ibUT)"
apply(subst nextM_IfTrue)
unfolding endPC_def unfolding prog_def by auto

lemma nextM_pc3_else[simp]: 
"vs xx < NN 
 nextM (Config 3 (State (Vstore vs) avst h p), ibT, ibUT) = 
 (Config 6 (State (Vstore vs) avst h p), ibT, ibUT)"
apply(subst nextM_IfFalse)
unfolding endPC_def unfolding prog_def by auto

lemma nextM_pc3: 
"nextM (Config 3 (State (Vstore vs) avst h p), ibT, ibUT) = 
 (Config (if vs xx < NN then 6 else 4) (State (Vstore vs) avst h p), ibT, ibUT)"
by(cases "vs xx < NN", auto)

lemma readLocs_pc3[simp]: 
"readLocs (Config 3 s) = {}"
unfolding endPC_def readLocs_def unfolding prog_def by auto

(* *)

lemma nextB_pc4[simp]: 
"nextB (Config 4 s, ibT, ibUT) = (Config 5 s, ibT, ibUT)"
apply(subst nextB_Start_Skip_Fence)
unfolding endPC_def unfolding prog_def by auto

lemma readLocs_pc4[simp]: 
"readLocs (Config 4 s) = {}"
unfolding endPC_def readLocs_def unfolding prog_def by auto

(* *)
lemma nextB_pc5[simp]:
"nextB (Config 5 (State (Vstore vs) avst (Heap h) p), ibT, ibUT) = 
(let l = (array_loc aa2 (nat (h (array_loc aa1 (nat (vs xx)) avst) * 512)) avst) 
  in (Config 6 (State (Vstore (vs(tt := h l))) avst (Heap h) p)), ibT, ibUT)"
apply(subst nextB_Assign)
unfolding endPC_def unfolding prog_def by auto


lemma readLocs_pc5[simp]: 
"readLocs (Config 5 (State (Vstore vs) avst (Heap h) p)) = 
 {array_loc aa2 (nat (h (array_loc aa1 (nat (vs xx)) avst) * 512)) avst, array_loc aa1 (nat (vs xx)) avst}"
unfolding endPC_def readLocs_def unfolding prog_def by auto

(**)

lemma nextB_pc6[simp]:
"nextB (Config 6 s, ibT, ibUT) = (Config 7 s, ibT, ibUT)"
apply(subst nextB_Output)
unfolding endPC_def unfolding prog_def by auto

lemma readLocs_pc6[simp]: 
"readLocs (Config 6 (State (Vstore vs) avst (Heap h) p)) = 
 {}"
unfolding endPC_def readLocs_def unfolding prog_def by auto
(* *)


lemma nextB_stepB_pc: 
"pc < 7  (pc = 1  ibUT  LNil)  
(Config pc s, ibT, ibUT) →B nextB (Config pc s, ibT, ibUT)"
apply(cases s) subgoal for vst avst hh p apply(cases vst, cases avst, cases hh)
subgoal for vs as h
using cases_6[of pc] apply safe
  subgoal by simp 
  subgoal by simp 
  (* *)
  subgoal apply simp apply(subst stepB.simps, unfold endPC_def)
  by (simp add: prog_def, metis llist.exhaust_sel) 
  subgoal apply simp apply(subst stepB.simps, unfold endPC_def)
    by (simp add: prog_def)  
  subgoal apply simp apply(subst stepB.simps, unfold endPC_def)
    by (simp add: prog_def)
  (* *)
  subgoal by(cases "vs xx < NN", simp_all)  
  subgoal by(cases "vs xx < NN", simp_all)  
  (* *)
  subgoal apply simp apply(subst stepB.simps) unfolding endPC_def 
  by (simp add: prog_def) 
  subgoal apply simp apply(subst stepB.simps) unfolding endPC_def 
  by (simp add: prog_def) 
  (* *)
  subgoal apply simp apply(subst stepB.simps) unfolding endPC_def 
  by (simp add: prog_def) 
  subgoal apply simp apply(subst stepB.simps) unfolding endPC_def 
  by (simp add: prog_def) 
  (* *)
  subgoal apply simp apply(subst stepB.simps) unfolding endPC_def 
  by (simp add: prog_def) 
  subgoal apply simp apply(subst stepB.simps) unfolding endPC_def 
    by (simp add: prog_def) 
  by simp+ . . 


lemma not_finalB: 
"pc < 7  (pc = 1  ibUT  LNil)  
 ¬ finalB (Config pc s, ibT, ibUT)"
using nextB_stepB_pc by (simp add: stepB_iff_nextB)

lemma finalB_pc_iff': 
"pc < 7 
 finalB (Config pc s, ibT, ibUT)  
 (pc = 1  ibUT = LNil)"
  subgoal apply safe
    subgoal using nextB_stepB_pc[of pc] by (auto simp add: stepB_iff_nextB) 
    subgoal using nextB_stepB_pc[of pc] by (auto simp add: stepB_iff_nextB) 
    subgoal using finalB_iff getUntrustedInput_pcOf by auto . .

lemma finalB_pc_iff: 
"pc  7 
 finalB (Config pc s, ibT, ibUT)  
 (pc = 1  ibUT = LNil  pc = 7)"
  using cases_6[of pc] apply (elim disjE, simp add: finalB_def)
  subgoal by (meson final_def stebB_0)
  by (simp add: finalB_pc_iff' finalB_endPC)+

lemma finalB_pcOf_iff[simp]:
"pcOf cfg  7  
 finalB (cfg, ibT, ibUT)  (pcOf cfg = 1  ibUT = LNil  pcOf cfg = 7)"
  by (metis config.collapse finalB_pc_iff) 

(*lemmas we need*)

lemma finalS_cond:"pcOf cfg < 7  cfgs = []  (pcOf cfg = 1  ibUT  LNil)  ¬ finalS (pstate, cfg, cfgs, ibT, ibUT, ls)"
  apply(cases cfg)
  subgoal for pc s apply(cases s)
  subgoal for vst avst hh p apply(cases vst, cases avst, cases hh)
    subgoal for vs as h
      using cases_6[of pc] apply(elim disjE) unfolding finalS_defs
      subgoal using nonspec_normal[of "[]" "Config pc (State (Vstore vs) avst hh p)" 
                                        pstate pstate ibT ibUT 
                                        "Config 1 (State (Vstore vs) avst hh p)" 
                                        ibT ibUT "[]" "ls  readLocs (Config pc (State (Vstore vs) avst hh p))" ls]
        using is_If_pc by force


      subgoal apply(frule nonspec_normal[of cfgs "Config pc (State (Vstore vs) avst hh p)" 
                                    pstate pstate ibT ibUT 
                                    "Config 2 (State (Vstore (vs(xx:= lhd ibUT))) avst hh p)" 
                                    ibT "ltl ibUT" "[]" "ls  readLocs (Config pc (State (Vstore vs) avst hh p))" ls])
              prefer 7 subgoal by metis by simp_all
      subgoal apply(frule nonspec_normal[of cfgs "Config pc (State (Vstore vs) avst hh p)" 
                                    pstate pstate ibT ibUT 
                                    "Config 3 (State (Vstore (vs(tt:= 0))) avst hh p)" 
                                    ibT ibUT "[]" "ls  readLocs (Config pc (State (Vstore vs) avst hh p))" ls])
              prefer 7 subgoal by metis by simp_all

    subgoal apply(cases "mispred pstate [3]")
      subgoal apply(frule nonspec_mispred[of cfgs "Config pc (State (Vstore vs) avst hh p)"
                                             pstate "update pstate [pcOf (Config pc (State (Vstore vs) avst hh p))]"
                                             ibT ibUT "Config (if vs xx < NN then 4 else 6) (State (Vstore vs) avst hh p)"
                                             ibT ibUT "Config (if vs xx < NN then 6 else 4) (State (Vstore vs) avst hh p)"
                                             ibT ibUT "[Config (if vs xx < NN then 6 else 4) (State (Vstore vs) avst hh p)]"
                                             "ls  readLocs (Config pc (State (Vstore vs) avst hh p))" ls])
      prefer 9 subgoal by metis by (simp add: finalM_iff)+

      subgoal apply(frule nonspec_normal[of cfgs "Config pc (State (Vstore vs) avst hh p)"
                                    pstate pstate ibT ibUT 
                                    "Config (if vs xx < NN then 4 else 6) (State (Vstore vs) avst hh p)" 
                                    ibT ibUT "[]" "ls  readLocs (Config pc (State (Vstore vs) avst hh p))" ls])
      prefer 7 subgoal by metis by simp_all . 

      subgoal apply(frule nonspec_normal[of cfgs "Config pc (State (Vstore vs) avst hh p)" 
                                    pstate pstate ibT ibUT 
                                    "Config 5 (State (Vstore vs) avst hh p)" 
                                    ibT ibUT "[]" ls ls])
              prefer 7 subgoal by metis by simp_all

    subgoal apply(frule nonspec_normal[of cfgs "Config pc (State (Vstore vs) avst hh p)" 
                         pstate pstate ibT ibUT 
                         "(let l = (array_loc aa2 (nat (h (array_loc aa1 (nat (vs xx)) avst) * 512)) avst) 
                                in (Config 6 (State (Vstore (vs(tt := h l))) avst hh p)))" 
                         ibT ibUT "[]" "ls  readLocs (Config pc (State (Vstore vs) avst hh p))" ls])
            prefer 7 subgoal by metis by simp_all

    subgoal apply(frule nonspec_normal[of cfgs "Config pc (State (Vstore vs) avst hh p)" 
                                    pstate pstate ibT ibUT 
                                    "Config 7 (State (Vstore vs) avst hh p)" 
                                    ibT ibUT "[]" ls ls])
            prefer 7 subgoal by metis by simp_all

    by simp_all . . .

lemma finalS_cond_spec:
      "pcOf cfg < 7  
       (pcOf (last cfgs) = 4  pcOf cfg = 6)  (pcOf (last cfgs) = 6  pcOf cfg = 4)  
       length cfgs =  Suc 0 
      ¬ finalS (pstate, cfg, cfgs, ibT, ibUT, ls)"
  apply(cases cfg)
  subgoal for pc s apply(cases s)
  subgoal for vst avst hh p apply(cases vst, cases avst, cases hh)
    subgoal for vs as h
      apply(elim disjE, elim conjE) unfolding finalS_defs
      subgoal using spec_resolve[of cfgs pstate cfg "update pstate (pcOf cfg # map pcOf cfgs)" cfg "[]" ibT ibT ibUT ibUT ls ls ]
              by (metis (no_types, lifting) butlast.simps(2) empty_set last_ConsL 
                  length_0_conv length_Suc_conv list.simps(8,9,15) pos2 resolve.simps)

      subgoal apply(elim conjE) 
              using spec_resolve[of cfgs pstate cfg "update pstate (pcOf cfg # map pcOf cfgs)" cfg "[]" ibT ibT ibUT ibUT ls ls ]
              by (metis (no_types, lifting) empty_set insert_commute last_ConsL resolve.simps
                  length_0_conv length_1_butlast length_Suc_conv list.simps(9,8,15)) . . . .


end