# HG changeset patch # User paulson # Date 982076573 -3600 # Node ID 83d03e966c687ff3849ce530409a5521d01d6099 # Parent ba314b436aab144678bddb4ae9eeff98b8bd371e partial conversion to Isar script style in HOL/Auth removes some .ML files diff -r ba314b436aab -r 83d03e966c68 src/HOL/Auth/Event_lemmas.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Auth/Event_lemmas.ML Tue Feb 13 16:02:53 2001 +0100 @@ -0,0 +1,233 @@ +(* Title: HOL/Auth/Event + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1996 University of Cambridge +*) + + +val knows_Cons = thm "knows_Cons"; +val used_Nil = thm "used_Nil"; +val used_Cons = thm "used_Cons"; + + +(*Inserted by default but later removed. This declaration lets the file + be re-loaded.*) +Addsimps [knows_Cons, used_Nil, used_Cons]; + +(**** Function "knows" ****) + +(** Simplifying parts (insert X (knows Spy evs)) + = parts {X} Un parts (knows Spy evs) -- since general case loops*) + +bind_thm ("parts_insert_knows_Spy", + parts_insert |> + read_instantiate_sg (sign_of thy) [("H", "knows Spy evs")]); + + +val expand_event_case = thm "event.split"; + +Goal "knows Spy (Says A B X # evs) = insert X (knows Spy evs)"; +by (Simp_tac 1); +qed "knows_Spy_Says"; + +(*The point of letting the Spy see "bad" agents' notes is to prevent + redundant case-splits on whether A=Spy and whether A:bad*) +Goal "knows Spy (Notes A X # evs) = \ +\ (if A:bad then insert X (knows Spy evs) else knows Spy evs)"; +by (Simp_tac 1); +qed "knows_Spy_Notes"; + +Goal "knows Spy (Gets A X # evs) = knows Spy evs"; +by (Simp_tac 1); +qed "knows_Spy_Gets"; + +Goal "knows Spy evs <= knows Spy (Says A B X # evs)"; +by (simp_tac (simpset() addsimps [subset_insertI]) 1); +qed "knows_Spy_subset_knows_Spy_Says"; + +Goal "knows Spy evs <= knows Spy (Notes A X # evs)"; +by (Simp_tac 1); +by (Fast_tac 1); +qed "knows_Spy_subset_knows_Spy_Notes"; + +Goal "knows Spy evs <= knows Spy (Gets A X # evs)"; +by (simp_tac (simpset() addsimps [subset_insertI]) 1); +qed "knows_Spy_subset_knows_Spy_Gets"; + +(*Spy sees what is sent on the traffic*) +Goal "Says A B X : set evs --> X : knows Spy evs"; +by (induct_tac "evs" 1); +by (ALLGOALS (asm_simp_tac (simpset() addsplits [expand_event_case]))); +qed_spec_mp "Says_imp_knows_Spy"; + +Goal "Notes A X : set evs --> A: bad --> X : knows Spy evs"; +by (induct_tac "evs" 1); +by (ALLGOALS (asm_simp_tac (simpset() addsplits [expand_event_case]))); +qed_spec_mp "Notes_imp_knows_Spy"; + + +(*Use with addSEs to derive contradictions from old Says events containing + items known to be fresh*) +bind_thms ("knows_Spy_partsEs", + make_elim (Says_imp_knows_Spy RS parts.Inj) :: partsEs); + +Addsimps [knows_Spy_Says, knows_Spy_Notes, knows_Spy_Gets]; + + +(*Begin lemmas about agents' knowledge*) + +Goal "knows A (Says A B X # evs) = insert X (knows A evs)"; +by (Simp_tac 1); +qed "knows_Says"; + +Goal "knows A (Notes A X # evs) = insert X (knows A evs)"; +by (Simp_tac 1); +qed "knows_Notes"; + +Goal "A ~= Spy --> knows A (Gets A X # evs) = insert X (knows A evs)"; +by (Simp_tac 1); +qed "knows_Gets"; + + +Goal "knows A evs <= knows A (Says A B X # evs)"; +by (simp_tac (simpset() addsimps [subset_insertI]) 1); +qed "knows_subset_knows_Says"; + +Goal "knows A evs <= knows A (Notes A X # evs)"; +by (simp_tac (simpset() addsimps [subset_insertI]) 1); +qed "knows_subset_knows_Notes"; + +Goal "knows A evs <= knows A (Gets A X # evs)"; +by (simp_tac (simpset() addsimps [subset_insertI]) 1); +qed "knows_subset_knows_Gets"; + +(*Agents know what they say*) +Goal "Says A B X : set evs --> X : knows A evs"; +by (induct_tac "evs" 1); +by (ALLGOALS (asm_simp_tac (simpset() addsplits [expand_event_case]))); +by (Blast_tac 1); +qed_spec_mp "Says_imp_knows"; + +(*Agents know what they note*) +Goal "Notes A X : set evs --> X : knows A evs"; +by (induct_tac "evs" 1); +by (ALLGOALS (asm_simp_tac (simpset() addsplits [expand_event_case]))); +by (Blast_tac 1); +qed_spec_mp "Notes_imp_knows"; + +(*Agents know what they receive*) +Goal "A ~= Spy --> Gets A X : set evs --> X : knows A evs"; +by (induct_tac "evs" 1); +by (ALLGOALS (asm_simp_tac (simpset() addsplits [expand_event_case]))); +qed_spec_mp "Gets_imp_knows_agents"; + + +(*What agents DIFFERENT FROM Spy know + was either said, or noted, or got, or known initially*) +Goal "[| X : knows A evs; A ~= Spy |] ==> EX B. \ +\ Says A B X : set evs | Gets A X : set evs | Notes A X : set evs | X : initState A"; +by (etac rev_mp 1); +by (induct_tac "evs" 1); +by (ALLGOALS (asm_simp_tac (simpset() addsplits [expand_event_case]))); +by (Blast_tac 1); +qed_spec_mp "knows_imp_Says_Gets_Notes_initState"; + +(*What the Spy knows -- for the time being -- + was either said or noted, or known initially*) +Goal "[| X : knows Spy evs |] ==> EX A B. \ +\ Says A B X : set evs | Notes A X : set evs | X : initState Spy"; +by (etac rev_mp 1); +by (induct_tac "evs" 1); +by (ALLGOALS (asm_simp_tac (simpset() addsplits [expand_event_case]))); +by (Blast_tac 1); +qed_spec_mp "knows_Spy_imp_Says_Notes_initState"; + +(*END lemmas about agents' knowledge*) + + + +(**** NOTE REMOVAL--laws above are cleaner, as they don't involve "case" ****) +Delsimps [knows_Cons]; + + +(*** Fresh nonces ***) + +Goal "parts (knows Spy evs) <= used evs"; +by (induct_tac "evs" 1); +by (ALLGOALS (asm_simp_tac + (simpset() addsimps [parts_insert_knows_Spy] + addsplits [expand_event_case]))); +by (ALLGOALS Blast_tac); +qed "parts_knows_Spy_subset_used"; + +bind_thm ("usedI", impOfSubs parts_knows_Spy_subset_used); +AddIs [usedI]; + +Goal "parts (initState B) <= used evs"; +by (induct_tac "evs" 1); +by (ALLGOALS (asm_simp_tac + (simpset() addsimps [parts_insert_knows_Spy] + addsplits [expand_event_case]))); +by (ALLGOALS Blast_tac); +bind_thm ("initState_into_used", impOfSubs (result())); + +Goal "used (Says A B X # evs) = parts{X} Un used evs"; +by (Simp_tac 1); +qed "used_Says"; +Addsimps [used_Says]; + +Goal "used (Notes A X # evs) = parts{X} Un used evs"; +by (Simp_tac 1); +qed "used_Notes"; +Addsimps [used_Notes]; + +Goal "used (Gets A X # evs) = used evs"; +by (Simp_tac 1); +qed "used_Gets"; +Addsimps [used_Gets]; + +Goal "used [] <= used evs"; +by (Simp_tac 1); +by (blast_tac (claset() addIs [initState_into_used]) 1); +qed "used_nil_subset"; + +(**** NOTE REMOVAL--laws above are cleaner, as they don't involve "case" ****) +Delsimps [used_Nil, used_Cons]; + + +(*For proving theorems of the form X ~: 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 analz.*) + +bind_thms ("analz_mono_contra", + [knows_Spy_subset_knows_Spy_Says RS analz_mono RS contra_subsetD, + knows_Spy_subset_knows_Spy_Notes RS analz_mono RS contra_subsetD, + knows_Spy_subset_knows_Spy_Gets RS analz_mono RS contra_subsetD]); + +val analz_mono_contra_tac = + let val analz_impI = inst "P" "?Y ~: analz (knows Spy ?evs)" impI + in + rtac analz_impI THEN' + REPEAT1 o + (dresolve_tac + [knows_Spy_subset_knows_Spy_Says RS analz_mono RS contra_subsetD, + knows_Spy_subset_knows_Spy_Notes RS analz_mono RS contra_subsetD, + knows_Spy_subset_knows_Spy_Gets RS analz_mono RS contra_subsetD]) + THEN' + mp_tac + end; + +fun Reception_tac i = + blast_tac (claset() addDs [Says_imp_knows_Spy RS parts.Inj, + impOfSubs (parts_insert RS equalityD1), + parts_trans, + Says_imp_knows_Spy RS analz.Inj, + impOfSubs analz_mono, analz_cut] + addIs [less_SucI]) i; + + +(*Compatibility for the old "spies" function*) +bind_thms ("spies_partsEs", knows_Spy_partsEs); +bind_thm ("Says_imp_spies", Says_imp_knows_Spy); +bind_thm ("parts_insert_spies", parts_insert_knows_Spy); diff -r ba314b436aab -r 83d03e966c68 src/HOL/Auth/Public_lemmas.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Auth/Public_lemmas.ML Tue Feb 13 16:02:53 2001 +0100 @@ -0,0 +1,180 @@ +(* Title: HOL/Auth/Public_lemmas + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1996 University of Cambridge + +Theory of Public Keys (common to all symmetric-key protocols) + +Server keys; initial states of agents; new nonces and keys; function "sees" +*) + +val inj_pubK = thm "inj_pubK"; +val priK_neq_pubK = thm "priK_neq_pubK"; + +(*** Basic properties of pubK & priK ***) + +AddIffs [inj_pubK RS inj_eq]; + +Goal "(priK A = priK B) = (A=B)"; +by Safe_tac; +by (dres_inst_tac [("f","invKey")] arg_cong 1); +by (Full_simp_tac 1); +qed "priK_inj_eq"; + +AddIffs [priK_inj_eq]; +AddIffs [priK_neq_pubK, priK_neq_pubK RS not_sym]; + +Goalw [isSymKey_def] "~ isSymKey (pubK A)"; +by (Simp_tac 1); +qed "not_isSymKey_pubK"; + +Goalw [isSymKey_def] "~ isSymKey (priK A)"; +by (Simp_tac 1); +qed "not_isSymKey_priK"; + +AddIffs [not_isSymKey_pubK, not_isSymKey_priK]; + + +(** "Image" equations that hold for injective functions **) + +Goal "(invKey x : invKey`A) = (x:A)"; +by Auto_tac; +qed "invKey_image_eq"; + +(*holds because invKey is injective*) +Goal "(pubK x : pubK`A) = (x:A)"; +by Auto_tac; +qed "pubK_image_eq"; + +Goal "(priK x ~: pubK`A)"; +by Auto_tac; +qed "priK_pubK_image_eq"; +Addsimps [invKey_image_eq, pubK_image_eq, priK_pubK_image_eq]; + + +(** Rewrites should not refer to initState(Friend i) + -- not in normal form! **) + +Goalw [keysFor_def] "keysFor (parts (initState C)) = {}"; +by (induct_tac "C" 1); +by (auto_tac (claset() addIs [range_eqI], simpset())); +qed "keysFor_parts_initState"; +Addsimps [keysFor_parts_initState]; + + +(*** Function "spies" ***) + +(*Agents see their own private keys!*) +Goal "Key (priK A) : initState A"; +by (induct_tac "A" 1); +by Auto_tac; +qed "priK_in_initState"; +AddIffs [priK_in_initState]; + +(*All public keys are visible*) +Goal "Key (pubK A) : spies evs"; +by (induct_tac "evs" 1); +by (ALLGOALS (asm_simp_tac + (simpset() addsimps [imageI, knows_Cons] + addsplits [expand_event_case]))); +qed_spec_mp "spies_pubK"; + +(*Spy sees private keys of bad agents!*) +Goal "A: bad ==> Key (priK A) : spies evs"; +by (induct_tac "evs" 1); +by (ALLGOALS (asm_simp_tac + (simpset() addsimps [imageI, knows_Cons] + addsplits [expand_event_case]))); +qed "Spy_spies_bad"; + +AddIffs [spies_pubK, spies_pubK RS analz.Inj]; +AddSIs [Spy_spies_bad]; + + +(*For not_bad_tac*) +Goal "[| Crypt (pubK A) X : analz (spies evs); A: bad |] \ +\ ==> X : analz (spies evs)"; +by (blast_tac (claset() addSDs [analz.Decrypt]) 1); +qed "Crypt_Spy_analz_bad"; + +(*Prove that the agent is uncompromised by the confidentiality of + a component of a message she's said.*) +fun not_bad_tac s = + case_tac ("(" ^ s ^ ") : bad") THEN' + SELECT_GOAL + (REPEAT_DETERM (dtac (Says_imp_spies RS analz.Inj) 1) THEN + REPEAT_DETERM (etac MPair_analz 1) THEN + THEN_BEST_FIRST + (dres_inst_tac [("A", s)] Crypt_Spy_analz_bad 1 THEN assume_tac 1) + (has_fewer_prems 1, size_of_thm) + Safe_tac); + + +(*** Fresh nonces ***) + +Goal "Nonce N ~: parts (initState B)"; +by (induct_tac "B" 1); +by Auto_tac; +qed "Nonce_notin_initState"; +AddIffs [Nonce_notin_initState]; + +Goal "Nonce N ~: used []"; +by (simp_tac (simpset() addsimps [used_Nil]) 1); +qed "Nonce_notin_used_empty"; +Addsimps [Nonce_notin_used_empty]; + + +(*** Supply fresh nonces for possibility theorems. ***) + +(*In any trace, there is an upper bound N on the greatest nonce in use.*) +Goal "EX N. ALL n. N<=n --> Nonce n ~: used evs"; +by (induct_tac "evs" 1); +by (res_inst_tac [("x","0")] exI 1); +by (ALLGOALS (asm_simp_tac + (simpset() addsimps [used_Cons] + addsplits [expand_event_case]))); +by Safe_tac; +by (ALLGOALS (rtac (msg_Nonce_supply RS exE))); +by (ALLGOALS (blast_tac (claset() addSEs [add_leE]))); +val lemma = result(); + +Goal "EX N. Nonce N ~: used evs"; +by (rtac (lemma RS exE) 1); +by (Blast_tac 1); +qed "Nonce_supply1"; + +Goal "Nonce (@ N. Nonce N ~: used evs) ~: used evs"; +by (rtac (lemma RS exE) 1); +by (rtac someI 1); +by (Fast_tac 1); +qed "Nonce_supply"; + +(*Tactic for possibility theorems*) +fun possibility_tac st = st |> + REPEAT (*omit used_Says so that Nonces start from different traces!*) + (ALLGOALS (simp_tac (simpset() delsimps [used_Says])) + THEN + REPEAT_FIRST (eq_assume_tac ORELSE' + resolve_tac [refl, conjI, Nonce_supply])); + + +(*** Specialized rewriting for the analz_image_... theorems ***) + +Goal "insert (Key K) H = Key ` {K} Un H"; +by (Blast_tac 1); +qed "insert_Key_singleton"; + +Goal "insert (Key K) (Key`KK Un C) = Key ` (insert K KK) Un C"; +by (Blast_tac 1); +qed "insert_Key_image"; + +(*Reverse the normal simplification of "image" to build up (not break down) + the set of keys. Based on analz_image_freshK_ss, but simpler.*) +val analz_image_keys_ss = + simpset() delsimps [image_insert, image_Un] + delsimps [imp_disjL] (*reduces blow-up*) + addsimps [image_insert RS sym, image_Un RS sym, + rangeI, + insert_Key_singleton, + insert_Key_image, Un_assoc RS sym]; + diff -r ba314b436aab -r 83d03e966c68 src/HOL/Auth/Shared_lemmas.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Auth/Shared_lemmas.ML Tue Feb 13 16:02:53 2001 +0100 @@ -0,0 +1,260 @@ +(* Title: HOL/Auth/Shared + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1996 University of Cambridge + +Theory of Shared Keys (common to all symmetric-key protocols) + +Shared, long-term keys; initial states of agents +*) + +val inj_shrK = thm "inj_shrK"; +val isSym_keys = thm "isSym_keys"; +val Key_supply_ax = thm "Key_supply_ax"; + +(*** Basic properties of shrK ***) + +(*Injectiveness: Agents' long-term keys are distinct.*) +AddIffs [inj_shrK RS inj_eq]; + +(* invKey(shrK A) = shrK A *) +Addsimps [rewrite_rule [isSymKey_def] isSym_keys]; + +(** Rewrites should not refer to initState(Friend i) + -- not in normal form! **) + +Goalw [keysFor_def] "keysFor (parts (initState C)) = {}"; +by (induct_tac "C" 1); +by Auto_tac; +qed "keysFor_parts_initState"; +Addsimps [keysFor_parts_initState]; + +(*Specialized to shared-key model: no need for addss in protocol proofs*) +Goal "[| K: keysFor (parts (insert X H)); X : synth (analz H) |] \ +\ ==> K : keysFor (parts H) | Key K : parts H"; +by (force_tac + (claset() addSDs [impOfSubs (parts_insert_subset_Un RS keysFor_mono), + impOfSubs (analz_subset_parts RS keysFor_mono)] + addIs [impOfSubs analz_subset_parts], + simpset()) 1); +qed "keysFor_parts_insert"; + +Goal "Crypt K X : H ==> K : keysFor H"; +by (dtac Crypt_imp_invKey_keysFor 1); +by (Asm_full_simp_tac 1); +qed "Crypt_imp_keysFor"; + + +(*** Function "knows" ***) + +(*Spy sees shared keys of agents!*) +Goal "A: bad ==> Key (shrK A) : knows Spy evs"; +by (induct_tac "evs" 1); +by (ALLGOALS (asm_simp_tac + (simpset() addsimps [imageI, knows_Cons] + addsplits [expand_event_case]))); +qed "Spy_knows_Spy_bad"; +AddSIs [Spy_knows_Spy_bad]; + +(*For not_bad_tac*) +Goal "[| Crypt (shrK A) X : analz (knows Spy evs); A: bad |] \ +\ ==> X : analz (knows Spy evs)"; +by (force_tac (claset() addSDs [analz.Decrypt], simpset()) 1); +qed "Crypt_Spy_analz_bad"; + +(*Prove that the agent is uncompromised by the confidentiality of + a component of a message she's said.*) +fun not_bad_tac s = + case_tac ("(" ^ s ^ ") : bad") THEN' + SELECT_GOAL + (REPEAT_DETERM (etac exE 1) THEN + REPEAT_DETERM (dtac (Says_imp_spies RS analz.Inj) 1) THEN + REPEAT_DETERM (etac MPair_analz 1) THEN + THEN_BEST_FIRST + (dres_inst_tac [("A", s)] Crypt_Spy_analz_bad 1 THEN assume_tac 1) + (has_fewer_prems 1, size_of_thm) + (Step_tac 1)); + + +(** Fresh keys never clash with long-term shared keys **) + +(*Agents see their own shared keys!*) +Goal "Key (shrK A) : initState A"; +by (induct_tac "A" 1); +by Auto_tac; +qed "shrK_in_initState"; +AddIffs [shrK_in_initState]; + +Goal "Key (shrK A) : used evs"; +by (rtac initState_into_used 1); +by (Blast_tac 1); +qed "shrK_in_used"; +AddIffs [shrK_in_used]; + +(*Used in parts_induct_tac and analz_Fake_tac to distinguish session keys + from long-term shared keys*) +Goal "Key K ~: used evs ==> K ~: range shrK"; +by (Blast_tac 1); +qed "Key_not_used"; + +Goal "Key K ~: used evs ==> shrK B ~= K"; +by (Blast_tac 1); +qed "shrK_neq"; + +Addsimps [Key_not_used, shrK_neq, shrK_neq RS not_sym]; + + +(*** Fresh nonces ***) + +Goal "Nonce N ~: parts (initState B)"; +by (induct_tac "B" 1); +by Auto_tac; +qed "Nonce_notin_initState"; +AddIffs [Nonce_notin_initState]; + +Goal "Nonce N ~: used []"; +by (simp_tac (simpset() addsimps [used_Nil]) 1); +qed "Nonce_notin_used_empty"; +Addsimps [Nonce_notin_used_empty]; + + +(*** Supply fresh nonces for possibility theorems. ***) + +(*In any trace, there is an upper bound N on the greatest nonce in use.*) +Goal "EX N. ALL n. N<=n --> Nonce n ~: used evs"; +by (induct_tac "evs" 1); +by (res_inst_tac [("x","0")] exI 1); +by (ALLGOALS (asm_simp_tac + (simpset() addsimps [used_Cons] + addsplits [expand_event_case]))); +by Safe_tac; +by (ALLGOALS (rtac (msg_Nonce_supply RS exE))); +by (ALLGOALS (blast_tac (claset() addSEs [add_leE]))); +val lemma = result(); + +Goal "EX N. Nonce N ~: used evs"; +by (rtac (lemma RS exE) 1); +by (Blast_tac 1); +qed "Nonce_supply1"; + +Goal "EX N N'. Nonce N ~: used evs & Nonce N' ~: used evs' & N ~= N'"; +by (cut_inst_tac [("evs","evs")] lemma 1); +by (cut_inst_tac [("evs","evs'")] lemma 1); +by (Clarify_tac 1); +by (res_inst_tac [("x","N")] exI 1); +by (res_inst_tac [("x","Suc (N+Na)")] exI 1); +by (asm_simp_tac (simpset() addsimps [less_not_refl3, le_add1, le_add2, + less_Suc_eq_le]) 1); +qed "Nonce_supply2"; + +Goal "EX N N' N''. Nonce N ~: used evs & Nonce N' ~: used evs' & \ +\ Nonce N'' ~: used evs'' & N ~= N' & N' ~= N'' & N ~= N''"; +by (cut_inst_tac [("evs","evs")] lemma 1); +by (cut_inst_tac [("evs","evs'")] lemma 1); +by (cut_inst_tac [("evs","evs''")] lemma 1); +by (Clarify_tac 1); +by (res_inst_tac [("x","N")] exI 1); +by (res_inst_tac [("x","Suc (N+Na)")] exI 1); +by (res_inst_tac [("x","Suc (Suc (N+Na+Nb))")] exI 1); +by (asm_simp_tac (simpset() addsimps [less_not_refl3, le_add1, le_add2, + less_Suc_eq_le]) 1); +qed "Nonce_supply3"; + +Goal "Nonce (@ N. Nonce N ~: used evs) ~: used evs"; +by (rtac (lemma RS exE) 1); +by (rtac someI 1); +by (Blast_tac 1); +qed "Nonce_supply"; + +(*** Supply fresh keys for possibility theorems. ***) + +Goal "EX K. Key K ~: used evs"; +by (rtac (Finites.emptyI RS Key_supply_ax RS exE) 1); +by (Blast_tac 1); +qed "Key_supply1"; + +Goal "EX K K'. Key K ~: used evs & Key K' ~: used evs' & K ~= K'"; +by (cut_inst_tac [("evs","evs")] (Finites.emptyI RS Key_supply_ax) 1); +by (etac exE 1); +by (cut_inst_tac [("evs","evs'")] + (Finites.emptyI RS Finites.insertI RS Key_supply_ax) 1); +by (Asm_full_simp_tac 1 THEN depth_tac (claset()) 2 1); (* replaces Auto_tac *) +qed "Key_supply2"; + +Goal "EX K K' K''. Key K ~: used evs & Key K' ~: used evs' & \ +\ Key K'' ~: used evs'' & K ~= K' & K' ~= K'' & K ~= K''"; +by (cut_inst_tac [("evs","evs")] (Finites.emptyI RS Key_supply_ax) 1); +by (etac exE 1); +(*Blast_tac requires instantiation of the keys for some reason*) +by (cut_inst_tac [("evs","evs'"), ("a1","K")] + (Finites.emptyI RS Finites.insertI RS Key_supply_ax) 1); +by (etac exE 1); +by (cut_inst_tac [("evs","evs''"), ("a1","Ka"), ("a2","K")] + (Finites.emptyI RS Finites.insertI RS Finites.insertI RS Key_supply_ax) 1); +by (Blast_tac 1); +qed "Key_supply3"; + +Goal "Key (@ K. Key K ~: used evs) ~: used evs"; +by (rtac (Finites.emptyI RS Key_supply_ax RS exE) 1); +by (rtac someI 1); +by (Blast_tac 1); +qed "Key_supply"; + +(*** Tactics for possibility theorems ***) + +(*Omitting used_Says makes the tactic much faster: it leaves expressions + such as Nonce ?N ~: used evs that match Nonce_supply*) +fun possibility_tac st = st |> + (REPEAT + (ALLGOALS (simp_tac (simpset() delsimps [used_Says, used_Notes, used_Gets] + setSolver safe_solver)) + THEN + REPEAT_FIRST (eq_assume_tac ORELSE' + resolve_tac [refl, conjI, Nonce_supply, Key_supply]))); + +(*For harder protocols (such as Recur) where we have to set up some + nonces and keys initially*) +fun basic_possibility_tac st = st |> + REPEAT + (ALLGOALS (asm_simp_tac (simpset() setSolver safe_solver)) + THEN + REPEAT_FIRST (resolve_tac [refl, conjI])); + + +(*** Specialized rewriting for analz_insert_freshK ***) + +Goal "A <= - (range shrK) ==> shrK x ~: A"; +by (Blast_tac 1); +qed "subset_Compl_range"; + +Goal "insert (Key K) H = Key ` {K} Un H"; +by (Blast_tac 1); +qed "insert_Key_singleton"; + +Goal "insert (Key K) (Key`KK Un C) = Key ` (insert K KK) Un C"; +by (Blast_tac 1); +qed "insert_Key_image"; + +(** Reverse the normal simplification of "image" to build up (not break down) + the set of keys. Use analz_insert_eq with (Un_upper2 RS analz_mono) to + erase occurrences of forwarded message components (X). **) + +bind_thms ("analz_image_freshK_simps", + simp_thms @ mem_simps @ (*these two allow its use with only:*) + disj_comms @ + [image_insert RS sym, image_Un RS sym, empty_subsetI, insert_subset, + analz_insert_eq, impOfSubs (Un_upper2 RS analz_mono), + insert_Key_singleton, subset_Compl_range, + Key_not_used, insert_Key_image, Un_assoc RS sym]); + +val analz_image_freshK_ss = + simpset() delsimps [image_insert, image_Un] + delsimps [imp_disjL] (*reduces blow-up*) + addsimps analz_image_freshK_simps; + +(*Lemma for the trivial direction of the if-and-only-if*) +Goal "(Key K : analz (Key`nE Un H)) --> (K : nE | Key K : analz H) ==> \ +\ (Key K : analz (Key`nE Un H)) = (K : nE | Key K : analz H)"; +by (blast_tac (claset() addIs [impOfSubs analz_mono]) 1); +qed "analz_image_freshK_lemma"; +