Theory Incomplete_fin

section ‹Relative Security Unwinding Incompleteness example›

text ‹Demonstrating a counterexample which is secure but fails in the finatary unwinding›

theory Incomplete_fin
  imports Unwinding_fin 
begin


no_notation bot ("")

(* This will be used for non-informative entities, e.g., a noninformative output: *)
abbreviation noninform ("") where "  undefined"


text ‹Demonstrating a counterexample which is secure but fails in the unwinding›
datatype State = ss | ss' 
type_synonym secret = "State"



fun transit::"State  State  bool"(infix "→I" 55) where 
  "transit s s' = (if (s = ss  s' = ss') then True else False)"

lemma transit_singlv:"ss →I ss'" by auto

lemma transit_iff:"s →I s'  (s = ss  s' = ss')" by auto

definition "final x  y. ¬ (→I) x y"

lemma final_sv'[simp]:"final ss'" unfolding final_def transit_iff by auto

lemma final_iff: "final s  s = ss'" unfolding final_def transit_iff by (auto,metis State.exhaust)

text ‹Vanilla-semantics system model›
type_synonym stateV = State
fun validTransV where "validTransV (s,s') = s →I s'"

text ‹No secrets or interaction›
fun isIntV :: "stateV  bool" where "isIntV s = False"
fun getIntV::"stateV  nat × nat" where "getIntV s = (,)"

definition isSecV :: "stateV  bool" where "isSecV s = False"
fun getSecV :: "stateV  secret" where "getSecV s = s"

(* *)

text ‹The optimization-enhanced system model›

type_synonym stateO = State
fun validTransO where "validTransO (s,s') = s →I s'"

text ‹No interaction, only isSec at starting state›
fun isIntO :: "stateO  bool" where "isIntO s = False"
fun getIntO::"stateO  nat × nat" where "getIntO s = (,)"


definition isSecO :: "stateO  bool" where "isSecO s = (if s = ss then True else False)"
fun getSecO :: "stateO  secret" where "getSecO s = s"

text ‹corrState›
fun corrState :: "stateV  stateO  bool" where 
"corrState cfgO cfgA = True"


interpretation Rel_Sec 
  where validTransV = validTransV and istateV = "λs. s = ss" 
  and finalV = final 
  and isSecV = isSecV and getSecV = getSecV
  and isIntV = isIntV and getIntV = getIntV
  (* *)
  and validTransO = validTransO and istateO = "λs. s = ss"
  and finalO = final
  and isSecO = isSecO  and getSecO = getSecO
  and isIntO = isIntO and getIntO = getIntV
  and corrState = corrState 
  apply(unfold_locales)
  subgoal by (auto simp: final_def)
  subgoal by (auto simp: final_def)
  subgoal by (auto simp: final_def isSecV_def)
  subgoal by (auto simp: final_def)
  subgoal by (auto simp: final_def)
  subgoal by (auto simp: final_iff isSecO_def) .


(*transition*)
lemma validFromV:"Van.validFromS ss [ss, ss']" unfolding Van.validFromS_def Van.validS_def by auto

lemma tr1_shape:"Opt.validFromS ss tr1  completedFromO ss tr1  tr1 = [ss, ss']"  
  unfolding completedFromO_def apply(erule disjE)
  subgoal by(simp add: final_iff)
  unfolding Opt.validFromS_def Opt.validS_def final_iff 
  apply(cases tr1, auto split: if_splits) 
    by (metis One_nat_def State.distinct append_Nil append_butlast_last_id diff_is_0_eq length_butlast length_greater_0_conv
        list.size(3) nth_Cons_0 nth_Cons_Suc verit_comp_simplify1(3))  

lemma tr1_shape':"s1 = ss  Opt.validFromS s1 tr1  completedFromO s1 tr1  tr1 = [ss, ss']"  
  using tr1_shape by auto

(*Systems are relatively secure*)

proposition "rsecure"
  unfolding rsecure_def3 
  apply(intro allI impI, elim conjE)
  apply(rule exI[of _ "ss"],rule exI[of _ "[ss, ss']"])
  apply(rule exI[of _ "ss"],rule exI[of _ "[ss, ss']"])
  by(frule tr1_shape', simp_all)+

lemma validSS:"Opt.validS [ss, ss']" unfolding Opt.validS_def by auto
lemma validSS_van:"Van.validS [ss, ss']" unfolding Van.validS_def by auto

lemma reachOs:"reachO ss" using Opt.reach.Istate by auto
lemma reachVs:"reachV ss" using Van.reach.Istate by auto
lemma reachO':"reachO ss'" using Opt.validS_reach[of "[ss, ss']", OF _ validSS] by auto
lemma reachV':"reachV ss'" using Van.validS_reach[of "[ss, ss']", OF _ validSS_van] by auto

lemma impE_eq:"x = x  Q  (Q  Rs)  Rs" by auto

lemma isSecOs:"isSecO ss" unfolding isSecO_def by auto
lemma neq_Sec:"¬eqSec ss ss" unfolding eqSec_def isSecO_def isSecV_def by auto

lemma statOs: "(sstatO' Eq ss ss) = Eq" unfolding sstatO'_def by auto

lemma noUnwind: shows init: "Δ  ss ss Eq ss ss Eq  unwindCond Δ  False"
  subgoal premises unwind using unwind(2)[unfolded unwindCond_def]
  (*step1*)
  apply-apply(erule allE[of _ ""])
  apply(erule allE[of _ "ss"], erule allE[of _ "ss"])
  apply(erule allE[of _ "Eq"])
  apply(erule allE[of _ "ss"], erule allE[of _ "ss"])
  apply(erule allE[of _ "Eq"], erule impE)
  subgoal using reachOs reachVs unwind(1) by auto
  apply(elim conjE disjE impE_eq)
  (*Proact*)
  subgoal apply(elim exE conjE, unfold proact_def, elim disjE)
    (************************)
    (*move 1 and next unwind*)
    (************************)
    subgoal for v unfolding move_1_def apply (simp split: if_splits)
      using unwind(2)[unfolded unwindCond_def]
      apply-apply(erule allE[of _ "v"])
      apply(erule allE[of _ "ss"], erule allE[of _ "ss"])
      apply(erule allE[of _ "Eq"])
      apply(erule allE[of _ "ss'"], erule allE[of _ "ss"])
      apply(erule allE[of _ "Eq"], erule impE)
      subgoal using reachOs reachVs reachV' unwind(1) by auto
      apply(elim conjE disjE impE_eq)
      (*Proact*)
      subgoal apply(elim exE conjE, unfold proact_def, elim disjE)
        (*move 1 not possible*)
        subgoal for v unfolding move_1_def by (simp split: if_splits) 
        (*Move 2 step*)
        subgoal for v' unfolding move_2_def apply (simp split: if_splits)
          using unwind(2)[unfolded unwindCond_def]
          apply-apply(erule allE[of _ "v'"])
          apply(erule allE[of _ "ss"], erule allE[of _ "ss"])
          apply(erule allE[of _ "Eq"])
          apply(erule allE[of _ "ss'"], erule allE[of _ "ss'"])
          apply(erule allE[of _ "Eq"], erule impE)
          subgoal using reachOs reachVs reachV' reachO' by auto
          apply(elim conjE disjE impE_eq)
          (*proact not possible*)    
          subgoal by(unfold proact_def, auto simp: move_1_def move_2_def move_12_def)
          (*React*)
          subgoal unfolding react_def match1_def match1_1_def match1_12_def 
              apply(elim conjE impE, simp)
            by(erule allE[of _ ss'], auto simp: neq_Sec isSecOs) .
        (*move12 not possible*)
        subgoal unfolding move_12_def by auto .
        (*react not possible*)
        unfolding react_def match1_def match1_1_def match1_12_def 
        apply(elim conjE impE, simp) by(erule allE[of _ ss'], auto simp: neq_Sec isSecOs) 

    (************************)
    (*move 2 and next unwind*)
    (************************)
    subgoal for v unfolding move_2_def apply (simp split: if_splits)
      using unwind(2)[unfolded unwindCond_def]
      apply-apply(erule allE[of _ "v"])
      apply(erule allE[of _ "ss"], erule allE[of _ "ss"])
      apply(erule allE[of _ "Eq"])
      apply(erule allE[of _ "ss"], erule allE[of _ "ss'"])
      apply(erule allE[of _ "Eq"], erule impE)
      subgoal using reachOs reachVs reachV' unwind(1) by auto
      apply(elim conjE disjE impE_eq)
      (*Proact*)
      subgoal apply(elim exE conjE, unfold proact_def, elim disjE)
        (*move 1 step*)
        subgoal for v' unfolding move_1_def apply (simp split: if_splits)
          using unwind(2)[unfolded unwindCond_def]
          apply-apply(erule allE[of _ "v'"])
          apply(erule allE[of _ "ss"], erule allE[of _ "ss"])
          apply(erule allE[of _ "Eq"])
          apply(erule allE[of _ "ss'"], erule allE[of _ "ss'"])
          apply(erule allE[of _ "Eq"], erule impE)
          subgoal using reachOs reachVs reachV' reachO' by auto
          apply(elim conjE disjE impE_eq)
          (*proact not possible*)    
          subgoal by(auto simp: proact_def move_1_def move_2_def move_12_def)
          (*React*)
          subgoal unfolding react_def match1_def match1_1_def match1_12_def 
              apply(elim conjE impE, simp)
            by(erule allE[of _ ss'], auto simp: neq_Sec isSecOs) .
        (*Move 2 not possible*)
        subgoal for v unfolding move_2_def by (simp split: if_splits) 
        (*move12 not possible*)
        subgoal unfolding move_12_def by auto .
        (*react not possible*)
        unfolding react_def match1_def match1_1_def match1_12_def 
        apply(elim conjE impE, simp) by(erule allE[of _ ss'], auto simp: neq_Sec isSecOs) 

    (************************)
    (*move 12*)
    (************************)
      subgoal for v unfolding move_12_def apply (simp split: if_splits)
      using unwind(2)[unfolded unwindCond_def]
      apply-apply(erule allE[of _ "v"])
      apply(erule allE[of _ "ss"], erule allE[of _ "ss"])
      apply(erule allE[of _ "Eq"])
      apply(erule allE[of _ "ss'"], erule allE[of _ "ss'"])
      apply(erule allE[of _ "Eq"], erule impE)
      subgoal using reachOs reachVs reachV' reachO' by (auto simp: statOs)
      apply(elim conjE disjE impE_eq)
      (*proact not possible*)    
      subgoal by(auto simp: proact_def move_1_def move_2_def move_12_def)
      (*react*)
      subgoal unfolding react_def match1_def match1_12_def match1_1_def apply(elim conjE impE, simp)
              by(erule allE[of _ ss'], auto simp: neq_Sec isSecOs) . .
  (*React*)
  subgoal unfolding react_def match1_def apply(elim conjE impE, simp)
    by(erule allE[of _ ss'], auto simp: neq_Sec isSecOs) . .


lemma incomplete_fin:
    assumes init: "initCond Δ"
    and unwind: "unwindCond Δ"
  shows "False"
  using noUnwind assms[unfolded initCond_def] by auto

end