| author | wenzelm | 
| Fri, 02 Jan 2009 19:30:12 +0100 | |
| changeset 29317 | 9faf1dfb4e7c | 
| parent 23746 | a455e69c31cc | 
| child 41775 | 6214816d79d3 | 
| permissions | -rw-r--r-- | 
| 17394 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 1 | (****************************************************************************** | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 2 | date: march 2002 | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 3 | author: Frederic Blanqui | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 4 | email: blanqui@lri.fr | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 5 | webpage: http://www.lri.fr/~blanqui/ | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 6 | |
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 7 | University of Cambridge, Computer Laboratory | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 8 | William Gates Building, JJ Thomson Avenue | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 9 | Cambridge CB3 0FD, United Kingdom | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 10 | ******************************************************************************) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 11 | |
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 12 | header{*Otway-Rees Protocol*}
 | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 13 | |
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 14 | theory Guard_OtwayRees imports Guard_Shared begin | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 15 | |
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 16 | subsection{*messages used in the protocol*}
 | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 17 | |
| 20768 | 18 | abbreviation | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20768diff
changeset | 19 | nil :: "msg" where | 
| 20768 | 20 | "nil == Number 0" | 
| 17394 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 21 | |
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20768diff
changeset | 22 | abbreviation | 
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20768diff
changeset | 23 | or1 :: "agent => agent => nat => event" where | 
| 20768 | 24 | "or1 A B NA == | 
| 25 |     Says A B {|Nonce NA, Agent A, Agent B, Ciph A {|Nonce NA, Agent A, Agent B|}|}"
 | |
| 17394 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 26 | |
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20768diff
changeset | 27 | abbreviation | 
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20768diff
changeset | 28 | or1' :: "agent => agent => agent => nat => msg => event" where | 
| 20768 | 29 |   "or1' A' A B NA X == Says A' B {|Nonce NA, Agent A, Agent B, X|}"
 | 
| 17394 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 30 | |
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20768diff
changeset | 31 | abbreviation | 
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20768diff
changeset | 32 | or2 :: "agent => agent => nat => nat => msg => event" where | 
| 20768 | 33 | "or2 A B NA NB X == | 
| 34 |     Says B Server {|Nonce NA, Agent A, Agent B, X,
 | |
| 17394 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 35 |                     Ciph B {|Nonce NA, Nonce NB, Agent A, Agent B|}|}"
 | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 36 | |
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20768diff
changeset | 37 | abbreviation | 
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20768diff
changeset | 38 | or2' :: "agent => agent => agent => nat => nat => event" where | 
| 20768 | 39 | "or2' B' A B NA NB == | 
| 40 |     Says B' Server {|Nonce NA, Agent A, Agent B,
 | |
| 17394 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 41 |                      Ciph A {|Nonce NA, Agent A, Agent B|},
 | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 42 |                      Ciph B {|Nonce NA, Nonce NB, Agent A, Agent B|}|}"
 | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 43 | |
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20768diff
changeset | 44 | abbreviation | 
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20768diff
changeset | 45 | or3 :: "agent => agent => nat => nat => key => event" where | 
| 20768 | 46 | "or3 A B NA NB K == | 
| 47 |     Says Server B {|Nonce NA, Ciph A {|Nonce NA, Key K|},
 | |
| 17394 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 48 |                     Ciph B {|Nonce NB, Key K|}|}"
 | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 49 | |
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20768diff
changeset | 50 | abbreviation | 
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20768diff
changeset | 51 | or3':: "agent => msg => agent => agent => nat => nat => key => event" where | 
| 20768 | 52 | "or3' S Y A B NA NB K == | 
| 53 |     Says S B {|Nonce NA, Y, Ciph B {|Nonce NB, Key K|}|}"
 | |
| 17394 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 54 | |
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20768diff
changeset | 55 | abbreviation | 
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20768diff
changeset | 56 | or4 :: "agent => agent => nat => msg => event" where | 
| 20768 | 57 |   "or4 A B NA X == Says B A {|Nonce NA, X, nil|}"
 | 
| 17394 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 58 | |
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20768diff
changeset | 59 | abbreviation | 
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20768diff
changeset | 60 | or4' :: "agent => agent => nat => key => event" where | 
| 20768 | 61 |   "or4' B' A NA K == Says B' A {|Nonce NA, Ciph A {|Nonce NA, Key K|}, nil|}"
 | 
| 17394 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 62 | |
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 63 | subsection{*definition of the protocol*}
 | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 64 | |
| 23746 | 65 | inductive_set or :: "event list set" | 
| 66 | where | |
| 17394 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 67 | |
| 23746 | 68 | Nil: "[]:or" | 
| 17394 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 69 | |
| 23746 | 70 | | Fake: "[| evs:or; X:synth (analz (spies evs)) |] ==> Says Spy B X # evs:or" | 
| 17394 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 71 | |
| 23746 | 72 | | OR1: "[| evs1:or; Nonce NA ~:used evs1 |] ==> or1 A B NA # evs1:or" | 
| 17394 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 73 | |
| 23746 | 74 | | OR2: "[| evs2:or; or1' A' A B NA X:set evs2; Nonce NB ~:used evs2 |] | 
| 75 | ==> or2 A B NA NB X # evs2:or" | |
| 17394 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 76 | |
| 23746 | 77 | | OR3: "[| evs3:or; or2' B' A B NA NB:set evs3; Key K ~:used evs3 |] | 
| 78 | ==> or3 A B NA NB K # evs3:or" | |
| 17394 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 79 | |
| 23746 | 80 | | OR4: "[| evs4:or; or2 A B NA NB X:set evs4; or3' S Y A B NA NB K:set evs4 |] | 
| 81 | ==> or4 A B NA X # evs4:or" | |
| 17394 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 82 | |
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 83 | subsection{*declarations for tactics*}
 | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 84 | |
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 85 | declare knows_Spy_partsEs [elim] | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 86 | declare Fake_parts_insert [THEN subsetD, dest] | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 87 | declare initState.simps [simp del] | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 88 | |
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 89 | subsection{*general properties of or*}
 | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 90 | |
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 91 | lemma or_has_no_Gets: "evs:or ==> ALL A X. Gets A X ~:set evs" | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 92 | by (erule or.induct, auto) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 93 | |
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 94 | lemma or_is_Gets_correct [iff]: "Gets_correct or" | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 95 | by (auto simp: Gets_correct_def dest: or_has_no_Gets) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 96 | |
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 97 | lemma or_is_one_step [iff]: "one_step or" | 
| 23746 | 98 | by (unfold one_step_def, clarify, ind_cases "ev#evs:or" for ev evs, auto) | 
| 17394 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 99 | |
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 100 | lemma or_has_only_Says' [rule_format]: "evs:or ==> | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 101 | ev:set evs --> (EX A B X. ev=Says A B X)" | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 102 | by (erule or.induct, auto) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 103 | |
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 104 | lemma or_has_only_Says [iff]: "has_only_Says or" | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 105 | by (auto simp: has_only_Says_def dest: or_has_only_Says') | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 106 | |
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 107 | subsection{*or is regular*}
 | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 108 | |
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 109 | lemma or1'_parts_spies [dest]: "or1' A' A B NA X:set evs | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 110 | ==> X:parts (spies evs)" | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 111 | by blast | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 112 | |
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 113 | lemma or2_parts_spies [dest]: "or2 A B NA NB X:set evs | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 114 | ==> X:parts (spies evs)" | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 115 | by blast | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 116 | |
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 117 | lemma or3_parts_spies [dest]: "Says S B {|NA, Y, Ciph B {|NB, K|}|}:set evs
 | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 118 | ==> K:parts (spies evs)" | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 119 | by blast | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 120 | |
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 121 | lemma or_is_regular [iff]: "regular or" | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 122 | apply (simp only: regular_def, clarify) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 123 | apply (erule or.induct, simp_all add: initState.simps knows.simps) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 124 | by (auto dest: parts_sub) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 125 | |
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 126 | subsection{*guardedness of KAB*}
 | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 127 | |
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 128 | lemma Guard_KAB [rule_format]: "[| evs:or; A ~:bad; B ~:bad |] ==> | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 129 | or3 A B NA NB K:set evs --> GuardK K {shrK A,shrK B} (spies evs)" 
 | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 130 | apply (erule or.induct) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 131 | (* Nil *) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 132 | apply simp_all | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 133 | (* Fake *) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 134 | apply (clarify, erule in_synth_GuardK, erule GuardK_analz, simp) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 135 | (* OR1 *) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 136 | apply blast | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 137 | (* OR2 *) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 138 | apply safe | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 139 | apply (blast dest: Says_imp_spies, blast) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 140 | (* OR3 *) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 141 | apply blast | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 142 | apply (drule_tac A=Server in Key_neq, simp+, rule No_Key, simp) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 143 | apply (drule_tac A=Server in Key_neq, simp+, rule No_Key, simp) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 144 | (* OR4 *) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 145 | by (blast dest: Says_imp_spies in_GuardK_kparts) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 146 | |
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 147 | subsection{*guardedness of NB*}
 | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 148 | |
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 149 | lemma Guard_NB [rule_format]: "[| evs:or; B ~:bad |] ==> | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 150 | or2 A B NA NB X:set evs --> Guard NB {shrK B} (spies evs)" 
 | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 151 | apply (erule or.induct) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 152 | (* Nil *) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 153 | apply simp_all | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 154 | (* Fake *) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 155 | apply safe | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 156 | apply (erule in_synth_Guard, erule Guard_analz, simp) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 157 | (* OR1 *) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 158 | apply (drule_tac n=NB in Nonce_neq, simp+, rule No_Nonce, simp) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 159 | apply (drule_tac n=NB in Nonce_neq, simp+, rule No_Nonce, simp) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 160 | (* OR2 *) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 161 | apply blast | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 162 | apply (drule_tac n=NA in Nonce_neq, simp+, rule No_Nonce, simp) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 163 | apply (blast intro!: No_Nonce dest: used_parts) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 164 | apply (drule_tac n=NA in Nonce_neq, simp+, rule No_Nonce, simp) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 165 | apply (blast intro!: No_Nonce dest: used_parts) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 166 | apply (blast dest: Says_imp_spies) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 167 | apply (blast dest: Says_imp_spies) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 168 | apply (case_tac "Ba=B", clarsimp) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 169 | apply (drule_tac n=NB and A=B in Nonce_neq, simp+) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 170 | apply (drule Says_imp_spies) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 171 | apply (drule_tac n'=NAa in in_Guard_kparts_neq, simp+, rule No_Nonce, simp) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 172 | (* OR3 *) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 173 | apply (drule Says_imp_spies) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 174 | apply (frule_tac n'=NAa in in_Guard_kparts_neq, simp+, rule No_Nonce, simp) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 175 | apply (case_tac "Aa=B", clarsimp) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 176 | apply (case_tac "NAa=NB", clarsimp) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 177 | apply (drule Says_imp_spies) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 178 | apply (drule_tac Y="{|Nonce NB, Agent Aa, Agent Ba|}"
 | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 179 | and K="shrK Aa" in in_Guard_kparts_Crypt, simp+) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 180 | apply (simp add: No_Nonce) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 181 | apply (case_tac "Ba=B", clarsimp) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 182 | apply (case_tac "NBa=NB", clarify) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 183 | apply (drule Says_imp_spies) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 184 | apply (drule_tac Y="{|Nonce NAa, Nonce NB, Agent Aa, Agent Ba|}"
 | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 185 | and K="shrK Ba" in in_Guard_kparts_Crypt, simp+) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 186 | apply (simp add: No_Nonce) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 187 | (* OR4 *) | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 188 | by (blast dest: Says_imp_spies)+ | 
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 189 | |
| 
a8c9ed3f9818
renamed Guard/NS_Public, Guard/OtwayRees, Guard/Yahalom.thy to avoid clash with plain Auth versions;
 wenzelm parents: diff
changeset | 190 | end |