| author | wenzelm | 
| Thu, 31 May 2001 20:53:49 +0200 | |
| changeset 11356 | 8fbb19b84f94 | 
| parent 11287 | 0103ee3082bf | 
| child 13922 | 75ae4244a596 | 
| permissions | -rw-r--r-- | 
| 2318 | 1 | (* Title: HOL/Auth/Public | 
| 2 | ID: $Id$ | |
| 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | |
| 4 | Copyright 1996 University of Cambridge | |
| 5 | ||
| 3512 
9dcb4daa15e8
Moving common declarations and proofs from theories "Shared"
 paulson parents: 
3478diff
changeset | 6 | Theory of Public Keys (common to all public-key protocols) | 
| 2318 | 7 | |
| 3512 
9dcb4daa15e8
Moving common declarations and proofs from theories "Shared"
 paulson parents: 
3478diff
changeset | 8 | Private and public keys; initial states of agents | 
| 2318 | 9 | *) | 
| 10 | ||
| 11104 | 11 | theory Public = Event | 
| 12 | files ("Public_lemmas.ML"):
 | |
| 2318 | 13 | |
| 14 | consts | |
| 11104 | 15 | pubK :: "agent => key" | 
| 2318 | 16 | |
| 17 | syntax | |
| 11104 | 18 | priK :: "agent => key" | 
| 2318 | 19 | |
| 3478 | 20 | translations (*BEWARE! expressions like (inj priK) will NOT work!*) | 
| 2318 | 21 | "priK x" == "invKey(pubK x)" | 
| 22 | ||
| 5183 | 23 | primrec | 
| 2318 | 24 | (*Agents know their private key and all public keys*) | 
| 11104 | 25 | initState_Server: "initState Server = | 
| 10833 | 26 | insert (Key (priK Server)) (Key ` range pubK)" | 
| 11104 | 27 | initState_Friend: "initState (Friend i) = | 
| 10833 | 28 | insert (Key (priK (Friend i))) (Key ` range pubK)" | 
| 11104 | 29 | initState_Spy: "initState Spy = | 
| 10833 | 30 | (Key`invKey`pubK`bad) Un (Key ` range pubK)" | 
| 2318 | 31 | |
| 32 | ||
| 11104 | 33 | axioms | 
| 2318 | 34 | (*Public keys are disjoint, and never clash with private keys*) | 
| 11104 | 35 | inj_pubK: "inj pubK" | 
| 36 | priK_neq_pubK: "priK A ~= pubK B" | |
| 37 | ||
| 38 | use "Public_lemmas.ML" | |
| 39 | ||
| 11287 | 40 | lemma [simp]: "K \\<in> symKeys ==> invKey K = K" | 
| 41 | by (simp add: symKeys_def) | |
| 42 | ||
| 11104 | 43 | (*Specialized methods*) | 
| 44 | ||
| 45 | method_setup possibility = {*
 | |
| 11270 
a315a3862bb4
better treatment of methods: uses Method.ctxt_args to refer to current
 paulson parents: 
11104diff
changeset | 46 | Method.ctxt_args (fn ctxt => | 
| 
a315a3862bb4
better treatment of methods: uses Method.ctxt_args to refer to current
 paulson parents: 
11104diff
changeset | 47 | Method.METHOD (fn facts => | 
| 
a315a3862bb4
better treatment of methods: uses Method.ctxt_args to refer to current
 paulson parents: 
11104diff
changeset | 48 | gen_possibility_tac (Simplifier.get_local_simpset ctxt))) *} | 
| 11104 | 49 | "for proving possibility theorems" | 
| 2318 | 50 | |
| 51 | end |