doc-src/TutorialI/Protocol/Public.thy
author berghofe
Mon, 30 Sep 2002 16:34:56 +0200
changeset 13604 57bfacbbaeda
parent 11250 c8bbf4c4bc2d
child 16417 9bc16273c2d4
permissions -rw-r--r--
- eliminated thin_leading_eqs_tac - gen_hyp_subst_tac now uses asm_lr_simp_tac instead of asm_full_simp_tac, in order to avoid divergence of new simplifier
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
11250
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
     1
(*  Title:      HOL/Auth/Public
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
     2
    ID:         $Id$
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
     4
    Copyright   1996  University of Cambridge
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
     5
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
     6
Theory of Public Keys (common to all public-key protocols)
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
     7
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
     8
Private and public keys; initial states of agents
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
     9
*)
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    10
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    11
theory Public = Event
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    12
files ("Public_lemmas.ML"):
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    13
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    14
consts
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    15
  pubK    :: "agent => key"
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    16
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    17
syntax
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    18
  priK    :: "agent => key"
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    19
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    20
translations  (*BEWARE! expressions like  (inj priK)  will NOT work!*)
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    21
  "priK x"  == "invKey(pubK x)"
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    22
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    23
primrec
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    24
        (*Agents know their private key and all public keys*)
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    25
  initState_Server:  "initState Server     =    
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    26
 		         insert (Key (priK Server)) (Key ` range pubK)"
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    27
  initState_Friend:  "initState (Friend i) =    
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    28
 		         insert (Key (priK (Friend i))) (Key ` range pubK)"
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    29
  initState_Spy:     "initState Spy        =    
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    30
 		         (Key`invKey`pubK`bad) Un (Key ` range pubK)"
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    31
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    32
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    33
axioms
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    34
  (*Public keys are disjoint, and never clash with private keys*)
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    35
  inj_pubK:        "inj pubK"
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    36
  priK_neq_pubK:   "priK A ~= pubK B"
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    37
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    38
use "Public_lemmas.ML"
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    39
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    40
(*Specialized methods*)
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    41
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    42
method_setup possibility = {*
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    43
    Method.no_args (Method.METHOD (fn facts => possibility_tac)) *}
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    44
    "for proving possibility theorems"
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    45
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    46
end