Theory Event

(*  Title:      HOL/Auth/Event.thy
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
    Copyright   1996  University of Cambridge

Datatype of events; function "spies"; freshness

"bad" agents have been broken by the Spy; their private keys and internal
    stores are visible to him
*)

section‹Theory of Events for Security Protocols›

theory Event imports Message begin

consts  ― ‹Initial states of agents --- a parameter of the construction›
  initState :: "agent  msg set"

datatype
  event = Says  agent agent msg
        | Gets  agent       msg
        | Notes agent       msg
       
consts 
  bad    :: "agent set"                         ― ‹compromised agents›

text‹Spy has access to his own key for spoof messages, but Server is secure›
specification (bad)
  Spy_in_bad     [iff]: "Spy  bad"
  Server_not_bad [iff]: "Server  bad"
    by (rule exI [of _ "{Spy}"], simp)

primrec knows :: "agent  event list  msg set"
where
  knows_Nil:   "knows A [] = initState A"
| knows_Cons:
    "knows A (ev # evs) =
       (if A = Spy then 
        (case ev of
           Says A' B X  insert X (knows Spy evs)
         | Gets A' X  knows Spy evs
         | Notes A' X   
             if A'  bad then insert X (knows Spy evs) else knows Spy evs)
        else
        (case ev of
           Says A' B X  
             if A'=A then insert X (knows A evs) else knows A evs
         | Gets A' X     
             if A'=A then insert X (knows A evs) else knows A evs
         | Notes A' X     
             if A'=A then insert X (knows A evs) else knows A evs))"
(*
  Case A=Spy on the Gets event
  enforces the fact that if a message is received then it must have been sent,
  therefore the oops case must use Notes
*)

text‹The constant "spies" is retained for compatibility's sake›

abbreviation (input)
  spies  :: "event list  msg set" where
  "spies  knows Spy"


text ‹Set of items that might be visible to somebody: complement of the set of fresh items›
primrec used :: "event list  msg set"
where
  used_Nil:   "used []         = (UN B. parts (initState B))"
| used_Cons:  "used (ev # evs) =
                     (case ev of
                        Says A B X  parts {X}  used evs
                      | Gets A X    used evs
                      | Notes A X   parts {X}  used evs)"
    ― ‹The case for termGets seems anomalous, but termGets always
        follows termSays in real protocols.  Seems difficult to change.
        See Gets_correct› in theory Guard/Extensions.thy›.›

lemma Notes_imp_used: "Notes A X  set evs  X  used evs"
  by (induct evs) (auto split: event.split) 

lemma Says_imp_used: "Says A B X  set evs  X  used evs"
  by (induct evs) (auto split: event.split) 


subsection‹Function termknows

(*Simplifying   
 parts(insert X (knows Spy evs)) = parts{X} ∪ parts(knows Spy evs).
  This version won't loop with the simplifier.*)
lemmas parts_insert_knows_A = parts_insert [of _ "knows A evs"] for A evs

lemma knows_Spy_Says [simp]:
  "knows Spy (Says A B X # evs) = insert X (knows Spy evs)"
  by simp

text‹Letting the Spy see "bad" agents' notes avoids redundant case-splits
      on whether termA=Spy and whether termAbad
lemma knows_Spy_Notes [simp]:
  "knows Spy (Notes A X # evs) =  
          (if Abad then insert X (knows Spy evs) else knows Spy evs)"
  by simp

lemma knows_Spy_Gets [simp]: "knows Spy (Gets A X # evs) = knows Spy evs"
  by simp

lemma knows_Spy_subset_knows_Spy_Says:
  "knows Spy evs  knows Spy (Says A B X # evs)"
  by (simp add: subset_insertI)

lemma knows_Spy_subset_knows_Spy_Notes:
  "knows Spy evs  knows Spy (Notes A X # evs)"
  by force

lemma knows_Spy_subset_knows_Spy_Gets:
  "knows Spy evs  knows Spy (Gets A X # evs)"
  by (simp add: subset_insertI)

text‹Spy sees what is sent on the traffic›
lemma Says_imp_knows_Spy:
     "Says A B X  set evs  X  knows Spy evs"
  by (induct evs) (auto split: event.split) 

lemma Notes_imp_knows_Spy [rule_format]:
     "Notes A X  set evs  A  bad  X  knows Spy evs"
  by (induct evs) (auto split: event.split) 


text‹Elimination rules: derive contradictions from old Says events containing
  items known to be fresh›
lemmas Says_imp_parts_knows_Spy = 
       Says_imp_knows_Spy [THEN parts.Inj, elim_format] 

lemmas knows_Spy_partsEs =
     Says_imp_parts_knows_Spy parts.Body [elim_format]

lemmas Says_imp_analz_Spy = Says_imp_knows_Spy [THEN analz.Inj]

text‹Compatibility for the old "spies" function›
lemmas spies_partsEs = knows_Spy_partsEs
lemmas Says_imp_spies = Says_imp_knows_Spy
lemmas parts_insert_spies = parts_insert_knows_A [of _ Spy]


subsection‹Knowledge of Agents›

lemma knows_subset_knows_Says: "knows A evs  knows A (Says A' B X # evs)"
  by (simp add: subset_insertI)

lemma knows_subset_knows_Notes: "knows A evs  knows A (Notes A' X # evs)"
  by (simp add: subset_insertI)

lemma knows_subset_knows_Gets: "knows A evs  knows A (Gets A' X # evs)"
  by (simp add: subset_insertI)

text‹Agents know what they say›
lemma Says_imp_knows [rule_format]: "Says A B X  set evs  X  knows A evs"
  by (induct evs) (force split: event.split)+

text‹Agents know what they note›
lemma Notes_imp_knows [rule_format]: "Notes A X  set evs  X  knows A evs"
  by (induct evs) (force split: event.split)+

text‹Agents know what they receive›
lemma Gets_imp_knows_agents [rule_format]:
     "A  Spy  Gets A X  set evs  X  knows A evs"
  by (induct evs) (force split: event.split)+

text‹What agents DIFFERENT FROM Spy know 
  was either said, or noted, or got, or known initially›
lemma knows_imp_Says_Gets_Notes_initState:
  "X  knows A evs; A  Spy  
     B. Says A B X  set evs  Gets A X  set evs  Notes A X  set evs  X  initState A"
  by(induct evs) (auto split: event.split_asm if_split_asm)

text‹What the Spy knows -- for the time being --
  was either said or noted, or known initially›
lemma knows_Spy_imp_Says_Notes_initState:
  "X  knows Spy evs  
    A B. Says A B X  set evs  Notes A X  set evs  X  initState Spy"
  by(induct evs) (auto split: event.split_asm if_split_asm)

lemma parts_knows_Spy_subset_used: "parts (knows Spy evs)  used evs"
  by (induct evs) (auto simp: parts_insert_knows_A split: event.split_asm if_split_asm)

lemmas usedI = parts_knows_Spy_subset_used [THEN subsetD, intro]

lemma initState_into_used: "X  parts (initState B)  X  used evs"
  by (induct evs) (auto simp: parts_insert_knows_A split: event.split)

text ‹New simprules to replace the primitive ones for @{term used} and @{term knows}

lemma used_Says [simp]: "used (Says A B X # evs) = parts{X}  used evs"
  by simp

lemma used_Notes [simp]: "used (Notes A X # evs) = parts{X}  used evs"
  by simp

lemma used_Gets [simp]: "used (Gets A X # evs) = used evs"
  by simp

lemma used_nil_subset: "used []  used evs"
  using initState_into_used by auto

text‹NOTE REMOVAL: the laws above are cleaner, as they don't involve "case"›
declare knows_Cons [simp del]
        used_Nil [simp del] used_Cons [simp del]


text‹For proving theorems of the form termX  analz (knows Spy evs)  P
  New events added by induction to "evs" are discarded.  Provided 
  this information isn't needed, the proof will be much shorter, since
  it will omit complicated reasoning about termanalz.›

lemmas analz_mono_contra =
       knows_Spy_subset_knows_Spy_Says [THEN analz_mono, THEN contra_subsetD]
       knows_Spy_subset_knows_Spy_Notes [THEN analz_mono, THEN contra_subsetD]
       knows_Spy_subset_knows_Spy_Gets [THEN analz_mono, THEN contra_subsetD]


lemma knows_subset_knows_Cons: "knows A evs  knows A (e # evs)"
  by (cases e, auto simp: knows_Cons)

lemma initState_subset_knows: "initState A  knows A evs"
  by (induct evs) (use knows_subset_knows_Cons in fastforce)+

text‹For proving new_keys_not_used›
lemma keysFor_parts_insert:
     "K  keysFor (parts (insert X G));  X  synth (analz H) 
       K  keysFor (parts (G  H)) | Key (invKey K)  parts H"
by (force 
    dest!: parts_insert_subset_Un [THEN keysFor_mono, THEN [2] rev_subsetD]
           analz_subset_parts [THEN keysFor_mono, THEN [2] rev_subsetD]
    intro: analz_subset_parts [THEN subsetD] parts_mono [THEN [2] rev_subsetD])


lemmas analz_impI = impI [where P = "Y  analz (knows Spy evs)"] for Y evs

ML
fun analz_mono_contra_tac ctxt = 
  resolve_tac ctxt @{thms analz_impI} THEN' 
  REPEAT1 o (dresolve_tac ctxt @{thms analz_mono_contra})
  THEN' (mp_tac ctxt)

method_setup analz_mono_contra = Scan.succeed (fn ctxt => SIMPLE_METHOD (REPEAT_FIRST (analz_mono_contra_tac ctxt)))
    "for proving theorems of the form X ∉ analz (knows Spy evs) ⟶ P"

text‹Useful for case analysis on whether a hash is a spoof or not›
lemmas syan_impI = impI [where P = "Y  synth (analz (knows Spy evs))"] for Y evs

ML
fun synth_analz_mono_contra_tac ctxt = 
  resolve_tac ctxt @{thms syan_impI} THEN'
  REPEAT1 o 
    (dresolve_tac ctxt
     [@{thm knows_Spy_subset_knows_Spy_Says} RS @{thm synth_analz_mono} RS @{thm contra_subsetD},
      @{thm knows_Spy_subset_knows_Spy_Notes} RS @{thm synth_analz_mono} RS @{thm contra_subsetD},
      @{thm knows_Spy_subset_knows_Spy_Gets} RS @{thm synth_analz_mono} RS @{thm contra_subsetD}])
  THEN'
  mp_tac ctxt

method_setup synth_analz_mono_contra = Scan.succeed (fn ctxt => SIMPLE_METHOD (REPEAT_FIRST (synth_analz_mono_contra_tac ctxt)))
    "for proving theorems of the form X ∉ synth (analz (knows Spy evs)) ⟶ P"

end